# HG changeset patch # User wenzelm # Date 1529848668 -7200 # Node ID eb53f944c8cd35ad189aebb7f2c535efe8cf84d3 # Parent 56034bd1b5f6a5825bb4ffd7872b7e9a3694b3e5 simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness; diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Bin.thy --- a/src/ZF/Bin.thy Sun Jun 24 11:41:41 2018 +0100 +++ b/src/ZF/Bin.thy Sun Jun 24 15:57:48 2018 +0200 @@ -16,7 +16,7 @@ section\Arithmetic on Binary Integers\ theory Bin -imports Int_ZF Datatype_ZF +imports Int Datatype begin consts bin :: i diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Cardinal.thy --- a/src/ZF/Cardinal.thy Sun Jun 24 11:41:41 2018 +0100 +++ b/src/ZF/Cardinal.thy Sun Jun 24 15:57:48 2018 +0200 @@ -5,7 +5,7 @@ section\Cardinal Numbers Without the Axiom of Choice\ -theory Cardinal imports OrderType Finite Nat_ZF Sum begin +theory Cardinal imports OrderType Finite Nat Sum begin definition (*least ordinal operator*) diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Datatype.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Datatype.thy Sun Jun 24 15:57:48 2018 +0200 @@ -0,0 +1,119 @@ +(* Title: ZF/Datatype.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1997 University of Cambridge +*) + +section\Datatype and CoDatatype Definitions\ + +theory Datatype +imports Inductive Univ QUniv +keywords "datatype" "codatatype" :: thy_decl +begin + +ML_file "Tools/datatype_package.ML" + +ML \ +(*Typechecking rules for most datatypes involving univ*) +structure Data_Arg = + struct + val intrs = + [@{thm SigmaI}, @{thm InlI}, @{thm InrI}, + @{thm Pair_in_univ}, @{thm Inl_in_univ}, @{thm Inr_in_univ}, + @{thm zero_in_univ}, @{thm A_into_univ}, @{thm nat_into_univ}, @{thm UnCI}]; + + + val elims = [make_elim @{thm InlD}, make_elim @{thm InrD}, (*for mutual recursion*) + @{thm SigmaE}, @{thm sumE}]; (*allows * and + in spec*) + end; + + +structure Data_Package = + Add_datatype_def_Fun + (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP + and Su=Standard_Sum + and Ind_Package = Ind_Package + and Datatype_Arg = Data_Arg + val coind = false); + + +(*Typechecking rules for most codatatypes involving quniv*) +structure CoData_Arg = + struct + val intrs = + [@{thm QSigmaI}, @{thm QInlI}, @{thm QInrI}, + @{thm QPair_in_quniv}, @{thm QInl_in_quniv}, @{thm QInr_in_quniv}, + @{thm zero_in_quniv}, @{thm A_into_quniv}, @{thm nat_into_quniv}, @{thm UnCI}]; + + val elims = [make_elim @{thm QInlD}, make_elim @{thm QInrD}, (*for mutual recursion*) + @{thm QSigmaE}, @{thm qsumE}]; (*allows * and + in spec*) + end; + +structure CoData_Package = + Add_datatype_def_Fun + (structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP + and Su=Quine_Sum + and Ind_Package = CoInd_Package + and Datatype_Arg = CoData_Arg + val coind = true); + + + +(*Simproc for freeness reasoning: compare datatype constructors for equality*) +structure DataFree = +struct + val trace = Unsynchronized.ref false; + + fun mk_new ([],[]) = Const(@{const_name True},FOLogic.oT) + | mk_new (largs,rargs) = + Balanced_Tree.make FOLogic.mk_conj + (map FOLogic.mk_eq (ListPair.zip (largs,rargs))); + + val datatype_ss = simpset_of @{context}; + + fun proc ctxt ct = + let val old = Thm.term_of ct + val thy = Proof_Context.theory_of ctxt + val _ = + if !trace then writeln ("data_free: OLD = " ^ Syntax.string_of_term ctxt old) + else () + val (lhs,rhs) = FOLogic.dest_eq old + val (lhead, largs) = strip_comb lhs + and (rhead, rargs) = strip_comb rhs + val lname = #1 (dest_Const lhead) handle TERM _ => raise Match; + val rname = #1 (dest_Const rhead) handle TERM _ => raise Match; + val lcon_info = the (Symtab.lookup (ConstructorsData.get thy) lname) + handle Option.Option => raise Match; + val rcon_info = the (Symtab.lookup (ConstructorsData.get thy) rname) + handle Option.Option => raise Match; + val new = + if #big_rec_name lcon_info = #big_rec_name rcon_info + andalso not (null (#free_iffs lcon_info)) then + if lname = rname then mk_new (largs, rargs) + else Const(@{const_name False},FOLogic.oT) + else raise Match + val _ = + if !trace then writeln ("NEW = " ^ Syntax.string_of_term ctxt new) + else (); + val goal = Logic.mk_equals (old, new) + val thm = Goal.prove ctxt [] [] goal + (fn _ => resolve_tac ctxt @{thms iff_reflection} 1 THEN + simp_tac (put_simpset datatype_ss ctxt addsimps #free_iffs lcon_info) 1) + handle ERROR msg => + (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Syntax.string_of_term ctxt goal); + raise Match) + in SOME thm end + handle Match => NONE; + + + val conv = + Simplifier.make_simproc @{context} "data_free" + {lhss = [@{term "(x::i) = y"}], proc = K proc}; + +end; +\ + +setup \ + Simplifier.map_theory_simpset (fn ctxt => ctxt addsimprocs [DataFree.conv]) +\ + +end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Datatype_ZF.thy --- a/src/ZF/Datatype_ZF.thy Sun Jun 24 11:41:41 2018 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,119 +0,0 @@ -(* Title: ZF/Datatype_ZF.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1997 University of Cambridge -*) - -section\Datatype and CoDatatype Definitions\ - -theory Datatype_ZF -imports Inductive_ZF Univ QUniv -keywords "datatype" "codatatype" :: thy_decl -begin - -ML_file "Tools/datatype_package.ML" - -ML \ -(*Typechecking rules for most datatypes involving univ*) -structure Data_Arg = - struct - val intrs = - [@{thm SigmaI}, @{thm InlI}, @{thm InrI}, - @{thm Pair_in_univ}, @{thm Inl_in_univ}, @{thm Inr_in_univ}, - @{thm zero_in_univ}, @{thm A_into_univ}, @{thm nat_into_univ}, @{thm UnCI}]; - - - val elims = [make_elim @{thm InlD}, make_elim @{thm InrD}, (*for mutual recursion*) - @{thm SigmaE}, @{thm sumE}]; (*allows * and + in spec*) - end; - - -structure Data_Package = - Add_datatype_def_Fun - (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP - and Su=Standard_Sum - and Ind_Package = Ind_Package - and Datatype_Arg = Data_Arg - val coind = false); - - -(*Typechecking rules for most codatatypes involving quniv*) -structure CoData_Arg = - struct - val intrs = - [@{thm QSigmaI}, @{thm QInlI}, @{thm QInrI}, - @{thm QPair_in_quniv}, @{thm QInl_in_quniv}, @{thm QInr_in_quniv}, - @{thm zero_in_quniv}, @{thm A_into_quniv}, @{thm nat_into_quniv}, @{thm UnCI}]; - - val elims = [make_elim @{thm QInlD}, make_elim @{thm QInrD}, (*for mutual recursion*) - @{thm QSigmaE}, @{thm qsumE}]; (*allows * and + in spec*) - end; - -structure CoData_Package = - Add_datatype_def_Fun - (structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP - and Su=Quine_Sum - and Ind_Package = CoInd_Package - and Datatype_Arg = CoData_Arg - val coind = true); - - - -(*Simproc for freeness reasoning: compare datatype constructors for equality*) -structure DataFree = -struct - val trace = Unsynchronized.ref false; - - fun mk_new ([],[]) = Const(@{const_name True},FOLogic.oT) - | mk_new (largs,rargs) = - Balanced_Tree.make FOLogic.mk_conj - (map FOLogic.mk_eq (ListPair.zip (largs,rargs))); - - val datatype_ss = simpset_of @{context}; - - fun proc ctxt ct = - let val old = Thm.term_of ct - val thy = Proof_Context.theory_of ctxt - val _ = - if !trace then writeln ("data_free: OLD = " ^ Syntax.string_of_term ctxt old) - else () - val (lhs,rhs) = FOLogic.dest_eq old - val (lhead, largs) = strip_comb lhs - and (rhead, rargs) = strip_comb rhs - val lname = #1 (dest_Const lhead) handle TERM _ => raise Match; - val rname = #1 (dest_Const rhead) handle TERM _ => raise Match; - val lcon_info = the (Symtab.lookup (ConstructorsData.get thy) lname) - handle Option.Option => raise Match; - val rcon_info = the (Symtab.lookup (ConstructorsData.get thy) rname) - handle Option.Option => raise Match; - val new = - if #big_rec_name lcon_info = #big_rec_name rcon_info - andalso not (null (#free_iffs lcon_info)) then - if lname = rname then mk_new (largs, rargs) - else Const(@{const_name False},FOLogic.oT) - else raise Match - val _ = - if !trace then writeln ("NEW = " ^ Syntax.string_of_term ctxt new) - else (); - val goal = Logic.mk_equals (old, new) - val thm = Goal.prove ctxt [] [] goal - (fn _ => resolve_tac ctxt @{thms iff_reflection} 1 THEN - simp_tac (put_simpset datatype_ss ctxt addsimps #free_iffs lcon_info) 1) - handle ERROR msg => - (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Syntax.string_of_term ctxt goal); - raise Match) - in SOME thm end - handle Match => NONE; - - - val conv = - Simplifier.make_simproc @{context} "data_free" - {lhss = [@{term "(x::i) = y"}], proc = K proc}; - -end; -\ - -setup \ - Simplifier.map_theory_simpset (fn ctxt => ctxt addsimprocs [DataFree.conv]) -\ - -end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Epsilon.thy --- a/src/ZF/Epsilon.thy Sun Jun 24 11:41:41 2018 +0100 +++ b/src/ZF/Epsilon.thy Sun Jun 24 15:57:48 2018 +0200 @@ -5,7 +5,7 @@ section\Epsilon Induction and Recursion\ -theory Epsilon imports Nat_ZF begin +theory Epsilon imports Nat begin definition eclose :: "i=>i" where diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Finite.thy --- a/src/ZF/Finite.thy Sun Jun 24 11:41:41 2018 +0100 +++ b/src/ZF/Finite.thy Sun Jun 24 15:57:48 2018 +0200 @@ -7,7 +7,7 @@ section\Finite Powerset Operator and Finite Function Space\ -theory Finite imports Inductive_ZF Epsilon Nat_ZF begin +theory Finite imports Inductive Epsilon Nat begin (*The natural numbers as a datatype*) rep_datatype diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Induct/Tree_Forest.thy --- a/src/ZF/Induct/Tree_Forest.thy Sun Jun 24 11:41:41 2018 +0100 +++ b/src/ZF/Induct/Tree_Forest.thy Sun Jun 24 15:57:48 2018 +0200 @@ -245,7 +245,7 @@ \ lemma preorder_map: - "z \ tree_forest(A) ==> preorder(map(h,z)) = List_ZF.map(h, preorder(z))" + "z \ tree_forest(A) ==> preorder(map(h,z)) = List.map(h, preorder(z))" by (induct set: tree_forest) (simp_all add: map_app_distrib) end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Inductive.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Inductive.thy Sun Jun 24 15:57:48 2018 +0200 @@ -0,0 +1,132 @@ +(* Title: ZF/Inductive.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Inductive definitions use least fixedpoints with standard products and sums +Coinductive definitions use greatest fixedpoints with Quine products and sums + +Sums are used only for mutual recursion; +Products are used only to derive "streamlined" induction rules for relations +*) + +section\Inductive and Coinductive Definitions\ + +theory Inductive +imports Fixedpt QPair Nat +keywords + "inductive" "coinductive" "inductive_cases" "rep_datatype" "primrec" :: thy_decl and + "domains" "intros" "monos" "con_defs" "type_intros" "type_elims" + "elimination" "induction" "case_eqns" "recursor_eqns" :: quasi_command +begin + +lemma def_swap_iff: "a == b ==> a = c \ c = b" + by blast + +lemma def_trans: "f == g ==> g(a) = b ==> f(a) = b" + by simp + +lemma refl_thin: "!!P. a = a ==> P ==> P" . + +ML_file "ind_syntax.ML" +ML_file "Tools/ind_cases.ML" +ML_file "Tools/cartprod.ML" +ML_file "Tools/inductive_package.ML" +ML_file "Tools/induct_tacs.ML" +ML_file "Tools/primrec_package.ML" + +ML \ +structure Lfp = + struct + val oper = @{const lfp} + val bnd_mono = @{const bnd_mono} + val bnd_monoI = @{thm bnd_monoI} + val subs = @{thm def_lfp_subset} + val Tarski = @{thm def_lfp_unfold} + val induct = @{thm def_induct} + end; + +structure Standard_Prod = + struct + val sigma = @{const Sigma} + val pair = @{const Pair} + val split_name = @{const_name split} + val pair_iff = @{thm Pair_iff} + val split_eq = @{thm split} + val fsplitI = @{thm splitI} + val fsplitD = @{thm splitD} + val fsplitE = @{thm splitE} + end; + +structure Standard_CP = CartProd_Fun (Standard_Prod); + +structure Standard_Sum = + struct + val sum = @{const sum} + val inl = @{const Inl} + val inr = @{const Inr} + val elim = @{const case} + val case_inl = @{thm case_Inl} + val case_inr = @{thm case_Inr} + val inl_iff = @{thm Inl_iff} + val inr_iff = @{thm Inr_iff} + val distinct = @{thm Inl_Inr_iff} + val distinct' = @{thm Inr_Inl_iff} + val free_SEs = Ind_Syntax.mk_free_SEs + [distinct, distinct', inl_iff, inr_iff, Standard_Prod.pair_iff] + end; + + +structure Ind_Package = + Add_inductive_def_Fun + (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP + and Su=Standard_Sum val coind = false); + + +structure Gfp = + struct + val oper = @{const gfp} + val bnd_mono = @{const bnd_mono} + val bnd_monoI = @{thm bnd_monoI} + val subs = @{thm def_gfp_subset} + val Tarski = @{thm def_gfp_unfold} + val induct = @{thm def_Collect_coinduct} + end; + +structure Quine_Prod = + struct + val sigma = @{const QSigma} + val pair = @{const QPair} + val split_name = @{const_name qsplit} + val pair_iff = @{thm QPair_iff} + val split_eq = @{thm qsplit} + val fsplitI = @{thm qsplitI} + val fsplitD = @{thm qsplitD} + val fsplitE = @{thm qsplitE} + end; + +structure Quine_CP = CartProd_Fun (Quine_Prod); + +structure Quine_Sum = + struct + val sum = @{const qsum} + val inl = @{const QInl} + val inr = @{const QInr} + val elim = @{const qcase} + val case_inl = @{thm qcase_QInl} + val case_inr = @{thm qcase_QInr} + val inl_iff = @{thm QInl_iff} + val inr_iff = @{thm QInr_iff} + val distinct = @{thm QInl_QInr_iff} + val distinct' = @{thm QInr_QInl_iff} + val free_SEs = Ind_Syntax.mk_free_SEs + [distinct, distinct', inl_iff, inr_iff, Quine_Prod.pair_iff] + end; + + +structure CoInd_Package = + Add_inductive_def_Fun(structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP + and Su=Quine_Sum val coind = true); + +\ + +end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Inductive_ZF.thy --- a/src/ZF/Inductive_ZF.thy Sun Jun 24 11:41:41 2018 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,132 +0,0 @@ -(* Title: ZF/Inductive_ZF.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1993 University of Cambridge - -Inductive definitions use least fixedpoints with standard products and sums -Coinductive definitions use greatest fixedpoints with Quine products and sums - -Sums are used only for mutual recursion; -Products are used only to derive "streamlined" induction rules for relations -*) - -section\Inductive and Coinductive Definitions\ - -theory Inductive_ZF -imports Fixedpt QPair Nat_ZF -keywords - "inductive" "coinductive" "inductive_cases" "rep_datatype" "primrec" :: thy_decl and - "domains" "intros" "monos" "con_defs" "type_intros" "type_elims" - "elimination" "induction" "case_eqns" "recursor_eqns" :: quasi_command -begin - -lemma def_swap_iff: "a == b ==> a = c \ c = b" - by blast - -lemma def_trans: "f == g ==> g(a) = b ==> f(a) = b" - by simp - -lemma refl_thin: "!!P. a = a ==> P ==> P" . - -ML_file "ind_syntax.ML" -ML_file "Tools/ind_cases.ML" -ML_file "Tools/cartprod.ML" -ML_file "Tools/inductive_package.ML" -ML_file "Tools/induct_tacs.ML" -ML_file "Tools/primrec_package.ML" - -ML \ -structure Lfp = - struct - val oper = @{const lfp} - val bnd_mono = @{const bnd_mono} - val bnd_monoI = @{thm bnd_monoI} - val subs = @{thm def_lfp_subset} - val Tarski = @{thm def_lfp_unfold} - val induct = @{thm def_induct} - end; - -structure Standard_Prod = - struct - val sigma = @{const Sigma} - val pair = @{const Pair} - val split_name = @{const_name split} - val pair_iff = @{thm Pair_iff} - val split_eq = @{thm split} - val fsplitI = @{thm splitI} - val fsplitD = @{thm splitD} - val fsplitE = @{thm splitE} - end; - -structure Standard_CP = CartProd_Fun (Standard_Prod); - -structure Standard_Sum = - struct - val sum = @{const sum} - val inl = @{const Inl} - val inr = @{const Inr} - val elim = @{const case} - val case_inl = @{thm case_Inl} - val case_inr = @{thm case_Inr} - val inl_iff = @{thm Inl_iff} - val inr_iff = @{thm Inr_iff} - val distinct = @{thm Inl_Inr_iff} - val distinct' = @{thm Inr_Inl_iff} - val free_SEs = Ind_Syntax.mk_free_SEs - [distinct, distinct', inl_iff, inr_iff, Standard_Prod.pair_iff] - end; - - -structure Ind_Package = - Add_inductive_def_Fun - (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP - and Su=Standard_Sum val coind = false); - - -structure Gfp = - struct - val oper = @{const gfp} - val bnd_mono = @{const bnd_mono} - val bnd_monoI = @{thm bnd_monoI} - val subs = @{thm def_gfp_subset} - val Tarski = @{thm def_gfp_unfold} - val induct = @{thm def_Collect_coinduct} - end; - -structure Quine_Prod = - struct - val sigma = @{const QSigma} - val pair = @{const QPair} - val split_name = @{const_name qsplit} - val pair_iff = @{thm QPair_iff} - val split_eq = @{thm qsplit} - val fsplitI = @{thm qsplitI} - val fsplitD = @{thm qsplitD} - val fsplitE = @{thm qsplitE} - end; - -structure Quine_CP = CartProd_Fun (Quine_Prod); - -structure Quine_Sum = - struct - val sum = @{const qsum} - val inl = @{const QInl} - val inr = @{const QInr} - val elim = @{const qcase} - val case_inl = @{thm qcase_QInl} - val case_inr = @{thm qcase_QInr} - val inl_iff = @{thm QInl_iff} - val inr_iff = @{thm QInr_iff} - val distinct = @{thm QInl_QInr_iff} - val distinct' = @{thm QInr_QInl_iff} - val free_SEs = Ind_Syntax.mk_free_SEs - [distinct, distinct', inl_iff, inr_iff, Quine_Prod.pair_iff] - end; - - -structure CoInd_Package = - Add_inductive_def_Fun(structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP - and Su=Quine_Sum val coind = true); - -\ - -end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/InfDatatype.thy --- a/src/ZF/InfDatatype.thy Sun Jun 24 11:41:41 2018 +0100 +++ b/src/ZF/InfDatatype.thy Sun Jun 24 15:57:48 2018 +0200 @@ -5,7 +5,7 @@ section\Infinite-Branching Datatype Definitions\ -theory InfDatatype imports Datatype_ZF Univ Finite Cardinal_AC begin +theory InfDatatype imports Datatype Univ Finite Cardinal_AC begin lemmas fun_Limit_VfromE = Limit_VfromE [OF apply_funtype InfCard_csucc [THEN InfCard_is_Limit]] diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Int.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Int.thy Sun Jun 24 15:57:48 2018 +0200 @@ -0,0 +1,921 @@ +(* Title: ZF/Int.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge +*) + +section\The Integers as Equivalence Classes Over Pairs of Natural Numbers\ + +theory Int imports EquivClass ArithSimp begin + +definition + intrel :: i where + "intrel == {p \ (nat*nat)*(nat*nat). + \x1 y1 x2 y2. p=<,> & x1#+y2 = x2#+y1}" + +definition + int :: i where + "int == (nat*nat)//intrel" + +definition + int_of :: "i=>i" \ \coercion from nat to int\ ("$# _" [80] 80) where + "$# m == intrel `` {}" + +definition + intify :: "i=>i" \ \coercion from ANYTHING to int\ where + "intify(m) == if m \ int then m else $#0" + +definition + raw_zminus :: "i=>i" where + "raw_zminus(z) == \\z. intrel``{}" + +definition + zminus :: "i=>i" ("$- _" [80] 80) where + "$- z == raw_zminus (intify(z))" + +definition + znegative :: "i=>o" where + "znegative(z) == \x y. xnat & \z" + +definition + iszero :: "i=>o" where + "iszero(z) == z = $# 0" + +definition + raw_nat_of :: "i=>i" where + "raw_nat_of(z) == natify (\\z. x#-y)" + +definition + nat_of :: "i=>i" where + "nat_of(z) == raw_nat_of (intify(z))" + +definition + zmagnitude :: "i=>i" where + \ \could be replaced by an absolute value function from int to int?\ + "zmagnitude(z) == + THE m. m\nat & ((~ znegative(z) & z = $# m) | + (znegative(z) & $- z = $# m))" + +definition + raw_zmult :: "[i,i]=>i" where + (*Cannot use UN here or in zadd because of the form of congruent2. + Perhaps a "curried" or even polymorphic congruent predicate would be + better.*) + "raw_zmult(z1,z2) == + \p1\z1. \p2\z2. split(%x1 y1. split(%x2 y2. + intrel``{}, p2), p1)" + +definition + zmult :: "[i,i]=>i" (infixl "$*" 70) where + "z1 $* z2 == raw_zmult (intify(z1),intify(z2))" + +definition + raw_zadd :: "[i,i]=>i" where + "raw_zadd (z1, z2) == + \z1\z1. \z2\z2. let =z1; =z2 + in intrel``{}" + +definition + zadd :: "[i,i]=>i" (infixl "$+" 65) where + "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))" + +definition + zdiff :: "[i,i]=>i" (infixl "$-" 65) where + "z1 $- z2 == z1 $+ zminus(z2)" + +definition + zless :: "[i,i]=>o" (infixl "$<" 50) where + "z1 $< z2 == znegative(z1 $- z2)" + +definition + zle :: "[i,i]=>o" (infixl "$\" 50) where + "z1 $\ z2 == z1 $< z2 | intify(z1)=intify(z2)" + + +declare quotientE [elim!] + +subsection\Proving that @{term intrel} is an equivalence relation\ + +(** Natural deduction for intrel **) + +lemma intrel_iff [simp]: + "<,>: intrel \ + x1\nat & y1\nat & x2\nat & y2\nat & x1#+y2 = x2#+y1" +by (simp add: intrel_def) + +lemma intrelI [intro!]: + "[| x1#+y2 = x2#+y1; x1\nat; y1\nat; x2\nat; y2\nat |] + ==> <,>: intrel" +by (simp add: intrel_def) + +lemma intrelE [elim!]: + "[| p \ intrel; + !!x1 y1 x2 y2. [| p = <,>; x1#+y2 = x2#+y1; + x1\nat; y1\nat; x2\nat; y2\nat |] ==> Q |] + ==> Q" +by (simp add: intrel_def, blast) + +lemma int_trans_lemma: + "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2 |] ==> x1 #+ y3 = x3 #+ y1" +apply (rule sym) +apply (erule add_left_cancel)+ +apply (simp_all (no_asm_simp)) +done + +lemma equiv_intrel: "equiv(nat*nat, intrel)" +apply (simp add: equiv_def refl_def sym_def trans_def) +apply (fast elim!: sym int_trans_lemma) +done + +lemma image_intrel_int: "[| m\nat; n\nat |] ==> intrel `` {} \ int" +by (simp add: int_def) + +declare equiv_intrel [THEN eq_equiv_class_iff, simp] +declare conj_cong [cong] + +lemmas eq_intrelD = eq_equiv_class [OF _ equiv_intrel] + +(** int_of: the injection from nat to int **) + +lemma int_of_type [simp,TC]: "$#m \ int" +by (simp add: int_def quotient_def int_of_def, auto) + +lemma int_of_eq [iff]: "($# m = $# n) \ natify(m)=natify(n)" +by (simp add: int_of_def) + +lemma int_of_inject: "[| $#m = $#n; m\nat; n\nat |] ==> m=n" +by (drule int_of_eq [THEN iffD1], auto) + + +(** intify: coercion from anything to int **) + +lemma intify_in_int [iff,TC]: "intify(x) \ int" +by (simp add: intify_def) + +lemma intify_ident [simp]: "n \ int ==> intify(n) = n" +by (simp add: intify_def) + + +subsection\Collapsing rules: to remove @{term intify} + from arithmetic expressions\ + +lemma intify_idem [simp]: "intify(intify(x)) = intify(x)" +by simp + +lemma int_of_natify [simp]: "$# (natify(m)) = $# m" +by (simp add: int_of_def) + +lemma zminus_intify [simp]: "$- (intify(m)) = $- m" +by (simp add: zminus_def) + +(** Addition **) + +lemma zadd_intify1 [simp]: "intify(x) $+ y = x $+ y" +by (simp add: zadd_def) + +lemma zadd_intify2 [simp]: "x $+ intify(y) = x $+ y" +by (simp add: zadd_def) + +(** Subtraction **) + +lemma zdiff_intify1 [simp]:"intify(x) $- y = x $- y" +by (simp add: zdiff_def) + +lemma zdiff_intify2 [simp]:"x $- intify(y) = x $- y" +by (simp add: zdiff_def) + +(** Multiplication **) + +lemma zmult_intify1 [simp]:"intify(x) $* y = x $* y" +by (simp add: zmult_def) + +lemma zmult_intify2 [simp]:"x $* intify(y) = x $* y" +by (simp add: zmult_def) + +(** Orderings **) + +lemma zless_intify1 [simp]:"intify(x) $< y \ x $< y" +by (simp add: zless_def) + +lemma zless_intify2 [simp]:"x $< intify(y) \ x $< y" +by (simp add: zless_def) + +lemma zle_intify1 [simp]:"intify(x) $\ y \ x $\ y" +by (simp add: zle_def) + +lemma zle_intify2 [simp]:"x $\ intify(y) \ x $\ y" +by (simp add: zle_def) + + +subsection\@{term zminus}: unary negation on @{term int}\ + +lemma zminus_congruent: "(%. intrel``{}) respects intrel" +by (auto simp add: congruent_def add_ac) + +lemma raw_zminus_type: "z \ int ==> raw_zminus(z) \ int" +apply (simp add: int_def raw_zminus_def) +apply (typecheck add: UN_equiv_class_type [OF equiv_intrel zminus_congruent]) +done + +lemma zminus_type [TC,iff]: "$-z \ int" +by (simp add: zminus_def raw_zminus_type) + +lemma raw_zminus_inject: + "[| raw_zminus(z) = raw_zminus(w); z \ int; w \ int |] ==> z=w" +apply (simp add: int_def raw_zminus_def) +apply (erule UN_equiv_class_inject [OF equiv_intrel zminus_congruent], safe) +apply (auto dest: eq_intrelD simp add: add_ac) +done + +lemma zminus_inject_intify [dest!]: "$-z = $-w ==> intify(z) = intify(w)" +apply (simp add: zminus_def) +apply (blast dest!: raw_zminus_inject) +done + +lemma zminus_inject: "[| $-z = $-w; z \ int; w \ int |] ==> z=w" +by auto + +lemma raw_zminus: + "[| x\nat; y\nat |] ==> raw_zminus(intrel``{}) = intrel `` {}" +apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent]) +done + +lemma zminus: + "[| x\nat; y\nat |] + ==> $- (intrel``{}) = intrel `` {}" +by (simp add: zminus_def raw_zminus image_intrel_int) + +lemma raw_zminus_zminus: "z \ int ==> raw_zminus (raw_zminus(z)) = z" +by (auto simp add: int_def raw_zminus) + +lemma zminus_zminus_intify [simp]: "$- ($- z) = intify(z)" +by (simp add: zminus_def raw_zminus_type raw_zminus_zminus) + +lemma zminus_int0 [simp]: "$- ($#0) = $#0" +by (simp add: int_of_def zminus) + +lemma zminus_zminus: "z \ int ==> $- ($- z) = z" +by simp + + +subsection\@{term znegative}: the test for negative integers\ + +lemma znegative: "[| x\nat; y\nat |] ==> znegative(intrel``{}) \ x natify(n)=0" +by (simp add: znegative int_of_def zminus Ord_0_lt_iff [THEN iff_sym]) + + +subsection\@{term nat_of}: Coercion of an Integer to a Natural Number\ + +lemma nat_of_intify [simp]: "nat_of(intify(z)) = nat_of(z)" +by (simp add: nat_of_def) + +lemma nat_of_congruent: "(\x. (\\x,y\. x #- y)(x)) respects intrel" +by (auto simp add: congruent_def split: nat_diff_split) + +lemma raw_nat_of: + "[| x\nat; y\nat |] ==> raw_nat_of(intrel``{}) = x#-y" +by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent]) + +lemma raw_nat_of_int_of: "raw_nat_of($# n) = natify(n)" +by (simp add: int_of_def raw_nat_of) + +lemma nat_of_int_of [simp]: "nat_of($# n) = natify(n)" +by (simp add: raw_nat_of_int_of nat_of_def) + +lemma raw_nat_of_type: "raw_nat_of(z) \ nat" +by (simp add: raw_nat_of_def) + +lemma nat_of_type [iff,TC]: "nat_of(z) \ nat" +by (simp add: nat_of_def raw_nat_of_type) + +subsection\zmagnitude: magnitide of an integer, as a natural number\ + +lemma zmagnitude_int_of [simp]: "zmagnitude($# n) = natify(n)" +by (auto simp add: zmagnitude_def int_of_eq) + +lemma natify_int_of_eq: "natify(x)=n ==> $#x = $# n" +apply (drule sym) +apply (simp (no_asm_simp) add: int_of_eq) +done + +lemma zmagnitude_zminus_int_of [simp]: "zmagnitude($- $# n) = natify(n)" +apply (simp add: zmagnitude_def) +apply (rule the_equality) +apply (auto dest!: not_znegative_imp_zero natify_int_of_eq + iff del: int_of_eq, auto) +done + +lemma zmagnitude_type [iff,TC]: "zmagnitude(z)\nat" +apply (simp add: zmagnitude_def) +apply (rule theI2, auto) +done + +lemma not_zneg_int_of: + "[| z \ int; ~ znegative(z) |] ==> \n\nat. z = $# n" +apply (auto simp add: int_def znegative int_of_def not_lt_iff_le) +apply (rename_tac x y) +apply (rule_tac x="x#-y" in bexI) +apply (auto simp add: add_diff_inverse2) +done + +lemma not_zneg_mag [simp]: + "[| z \ int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z" +by (drule not_zneg_int_of, auto) + +lemma zneg_int_of: + "[| znegative(z); z \ int |] ==> \n\nat. z = $- ($# succ(n))" +by (auto simp add: int_def znegative zminus int_of_def dest!: less_imp_succ_add) + +lemma zneg_mag [simp]: + "[| znegative(z); z \ int |] ==> $# (zmagnitude(z)) = $- z" +by (drule zneg_int_of, auto) + +lemma int_cases: "z \ int ==> \n\nat. z = $# n | z = $- ($# succ(n))" +apply (case_tac "znegative (z) ") +prefer 2 apply (blast dest: not_zneg_mag sym) +apply (blast dest: zneg_int_of) +done + +lemma not_zneg_raw_nat_of: + "[| ~ znegative(z); z \ int |] ==> $# (raw_nat_of(z)) = z" +apply (drule not_zneg_int_of) +apply (auto simp add: raw_nat_of_type raw_nat_of_int_of) +done + +lemma not_zneg_nat_of_intify: + "~ znegative(intify(z)) ==> $# (nat_of(z)) = intify(z)" +by (simp (no_asm_simp) add: nat_of_def not_zneg_raw_nat_of) + +lemma not_zneg_nat_of: "[| ~ znegative(z); z \ int |] ==> $# (nat_of(z)) = z" +apply (simp (no_asm_simp) add: not_zneg_nat_of_intify) +done + +lemma zneg_nat_of [simp]: "znegative(intify(z)) ==> nat_of(z) = 0" +apply (subgoal_tac "intify(z) \ int") +apply (simp add: int_def) +apply (auto simp add: znegative nat_of_def raw_nat_of + split: nat_diff_split) +done + + +subsection\@{term zadd}: addition on int\ + +text\Congruence Property for Addition\ +lemma zadd_congruent2: + "(%z1 z2. let =z1; =z2 + in intrel``{}) + respects2 intrel" +apply (simp add: congruent2_def) +(*Proof via congruent2_commuteI seems longer*) +apply safe +apply (simp (no_asm_simp) add: add_assoc Let_def) +(*The rest should be trivial, but rearranging terms is hard + add_ac does not help rewriting with the assumptions.*) +apply (rule_tac m1 = x1a in add_left_commute [THEN ssubst]) +apply (rule_tac m1 = x2a in add_left_commute [THEN ssubst]) +apply (simp (no_asm_simp) add: add_assoc [symmetric]) +done + +lemma raw_zadd_type: "[| z \ int; w \ int |] ==> raw_zadd(z,w) \ int" +apply (simp add: int_def raw_zadd_def) +apply (rule UN_equiv_class_type2 [OF equiv_intrel zadd_congruent2], assumption+) +apply (simp add: Let_def) +done + +lemma zadd_type [iff,TC]: "z $+ w \ int" +by (simp add: zadd_def raw_zadd_type) + +lemma raw_zadd: + "[| x1\nat; y1\nat; x2\nat; y2\nat |] + ==> raw_zadd (intrel``{}, intrel``{}) = + intrel `` {}" +apply (simp add: raw_zadd_def + UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2]) +apply (simp add: Let_def) +done + +lemma zadd: + "[| x1\nat; y1\nat; x2\nat; y2\nat |] + ==> (intrel``{}) $+ (intrel``{}) = + intrel `` {}" +by (simp add: zadd_def raw_zadd image_intrel_int) + +lemma raw_zadd_int0: "z \ int ==> raw_zadd ($#0,z) = z" +by (auto simp add: int_def int_of_def raw_zadd) + +lemma zadd_int0_intify [simp]: "$#0 $+ z = intify(z)" +by (simp add: zadd_def raw_zadd_int0) + +lemma zadd_int0: "z \ int ==> $#0 $+ z = z" +by simp + +lemma raw_zminus_zadd_distrib: + "[| z \ int; w \ int |] ==> $- raw_zadd(z,w) = raw_zadd($- z, $- w)" +by (auto simp add: zminus raw_zadd int_def) + +lemma zminus_zadd_distrib [simp]: "$- (z $+ w) = $- z $+ $- w" +by (simp add: zadd_def raw_zminus_zadd_distrib) + +lemma raw_zadd_commute: + "[| z \ int; w \ int |] ==> raw_zadd(z,w) = raw_zadd(w,z)" +by (auto simp add: raw_zadd add_ac int_def) + +lemma zadd_commute: "z $+ w = w $+ z" +by (simp add: zadd_def raw_zadd_commute) + +lemma raw_zadd_assoc: + "[| z1: int; z2: int; z3: int |] + ==> raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))" +by (auto simp add: int_def raw_zadd add_assoc) + +lemma zadd_assoc: "(z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)" +by (simp add: zadd_def raw_zadd_type raw_zadd_assoc) + +(*For AC rewriting*) +lemma zadd_left_commute: "z1$+(z2$+z3) = z2$+(z1$+z3)" +apply (simp add: zadd_assoc [symmetric]) +apply (simp add: zadd_commute) +done + +(*Integer addition is an AC operator*) +lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute + +lemma int_of_add: "$# (m #+ n) = ($#m) $+ ($#n)" +by (simp add: int_of_def zadd) + +lemma int_succ_int_1: "$# succ(m) = $# 1 $+ ($# m)" +by (simp add: int_of_add [symmetric] natify_succ) + +lemma int_of_diff: + "[| m\nat; n \ m |] ==> $# (m #- n) = ($#m) $- ($#n)" +apply (simp add: int_of_def zdiff_def) +apply (frule lt_nat_in_nat) +apply (simp_all add: zadd zminus add_diff_inverse2) +done + +lemma raw_zadd_zminus_inverse: "z \ int ==> raw_zadd (z, $- z) = $#0" +by (auto simp add: int_def int_of_def zminus raw_zadd add_commute) + +lemma zadd_zminus_inverse [simp]: "z $+ ($- z) = $#0" +apply (simp add: zadd_def) +apply (subst zminus_intify [symmetric]) +apply (rule intify_in_int [THEN raw_zadd_zminus_inverse]) +done + +lemma zadd_zminus_inverse2 [simp]: "($- z) $+ z = $#0" +by (simp add: zadd_commute zadd_zminus_inverse) + +lemma zadd_int0_right_intify [simp]: "z $+ $#0 = intify(z)" +by (rule trans [OF zadd_commute zadd_int0_intify]) + +lemma zadd_int0_right: "z \ int ==> z $+ $#0 = z" +by simp + + +subsection\@{term zmult}: Integer Multiplication\ + +text\Congruence property for multiplication\ +lemma zmult_congruent2: + "(%p1 p2. split(%x1 y1. split(%x2 y2. + intrel``{}, p2), p1)) + respects2 intrel" +apply (rule equiv_intrel [THEN congruent2_commuteI], auto) +(*Proof that zmult is congruent in one argument*) +apply (rename_tac x y) +apply (frule_tac t = "%u. x#*u" in sym [THEN subst_context]) +apply (drule_tac t = "%u. y#*u" in subst_context) +apply (erule add_left_cancel)+ +apply (simp_all add: add_mult_distrib_left) +done + + +lemma raw_zmult_type: "[| z \ int; w \ int |] ==> raw_zmult(z,w) \ int" +apply (simp add: int_def raw_zmult_def) +apply (rule UN_equiv_class_type2 [OF equiv_intrel zmult_congruent2], assumption+) +apply (simp add: Let_def) +done + +lemma zmult_type [iff,TC]: "z $* w \ int" +by (simp add: zmult_def raw_zmult_type) + +lemma raw_zmult: + "[| x1\nat; y1\nat; x2\nat; y2\nat |] + ==> raw_zmult(intrel``{}, intrel``{}) = + intrel `` {}" +by (simp add: raw_zmult_def + UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2]) + +lemma zmult: + "[| x1\nat; y1\nat; x2\nat; y2\nat |] + ==> (intrel``{}) $* (intrel``{}) = + intrel `` {}" +by (simp add: zmult_def raw_zmult image_intrel_int) + +lemma raw_zmult_int0: "z \ int ==> raw_zmult ($#0,z) = $#0" +by (auto simp add: int_def int_of_def raw_zmult) + +lemma zmult_int0 [simp]: "$#0 $* z = $#0" +by (simp add: zmult_def raw_zmult_int0) + +lemma raw_zmult_int1: "z \ int ==> raw_zmult ($#1,z) = z" +by (auto simp add: int_def int_of_def raw_zmult) + +lemma zmult_int1_intify [simp]: "$#1 $* z = intify(z)" +by (simp add: zmult_def raw_zmult_int1) + +lemma zmult_int1: "z \ int ==> $#1 $* z = z" +by simp + +lemma raw_zmult_commute: + "[| z \ int; w \ int |] ==> raw_zmult(z,w) = raw_zmult(w,z)" +by (auto simp add: int_def raw_zmult add_ac mult_ac) + +lemma zmult_commute: "z $* w = w $* z" +by (simp add: zmult_def raw_zmult_commute) + +lemma raw_zmult_zminus: + "[| z \ int; w \ int |] ==> raw_zmult($- z, w) = $- raw_zmult(z, w)" +by (auto simp add: int_def zminus raw_zmult add_ac) + +lemma zmult_zminus [simp]: "($- z) $* w = $- (z $* w)" +apply (simp add: zmult_def raw_zmult_zminus) +apply (subst zminus_intify [symmetric], rule raw_zmult_zminus, auto) +done + +lemma zmult_zminus_right [simp]: "w $* ($- z) = $- (w $* z)" +by (simp add: zmult_commute [of w]) + +lemma raw_zmult_assoc: + "[| z1: int; z2: int; z3: int |] + ==> raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))" +by (auto simp add: int_def raw_zmult add_mult_distrib_left add_ac mult_ac) + +lemma zmult_assoc: "(z1 $* z2) $* z3 = z1 $* (z2 $* z3)" +by (simp add: zmult_def raw_zmult_type raw_zmult_assoc) + +(*For AC rewriting*) +lemma zmult_left_commute: "z1$*(z2$*z3) = z2$*(z1$*z3)" +apply (simp add: zmult_assoc [symmetric]) +apply (simp add: zmult_commute) +done + +(*Integer multiplication is an AC operator*) +lemmas zmult_ac = zmult_assoc zmult_commute zmult_left_commute + +lemma raw_zadd_zmult_distrib: + "[| z1: int; z2: int; w \ int |] + ==> raw_zmult(raw_zadd(z1,z2), w) = + raw_zadd (raw_zmult(z1,w), raw_zmult(z2,w))" +by (auto simp add: int_def raw_zadd raw_zmult add_mult_distrib_left add_ac mult_ac) + +lemma zadd_zmult_distrib: "(z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)" +by (simp add: zmult_def zadd_def raw_zadd_type raw_zmult_type + raw_zadd_zmult_distrib) + +lemma zadd_zmult_distrib2: "w $* (z1 $+ z2) = (w $* z1) $+ (w $* z2)" +by (simp add: zmult_commute [of w] zadd_zmult_distrib) + +lemmas int_typechecks = + int_of_type zminus_type zmagnitude_type zadd_type zmult_type + + +(*** Subtraction laws ***) + +lemma zdiff_type [iff,TC]: "z $- w \ int" +by (simp add: zdiff_def) + +lemma zminus_zdiff_eq [simp]: "$- (z $- y) = y $- z" +by (simp add: zdiff_def zadd_commute) + +lemma zdiff_zmult_distrib: "(z1 $- z2) $* w = (z1 $* w) $- (z2 $* w)" +apply (simp add: zdiff_def) +apply (subst zadd_zmult_distrib) +apply (simp add: zmult_zminus) +done + +lemma zdiff_zmult_distrib2: "w $* (z1 $- z2) = (w $* z1) $- (w $* z2)" +by (simp add: zmult_commute [of w] zdiff_zmult_distrib) + +lemma zadd_zdiff_eq: "x $+ (y $- z) = (x $+ y) $- z" +by (simp add: zdiff_def zadd_ac) + +lemma zdiff_zadd_eq: "(x $- y) $+ z = (x $+ z) $- y" +by (simp add: zdiff_def zadd_ac) + + +subsection\The "Less Than" Relation\ + +(*"Less than" is a linear ordering*) +lemma zless_linear_lemma: + "[| z \ int; w \ int |] ==> z$ int; y \ int |] ==> (x \ y) \ (x $< y | y $< x)" +by (cut_tac z = x and w = y in zless_linear, auto) + +lemma zless_imp_intify_neq: "w $< z ==> intify(w) \ intify(z)" +apply auto +apply (subgoal_tac "~ (intify (w) $< intify (z))") +apply (erule_tac [2] ssubst) +apply (simp (no_asm_use)) +apply auto +done + +(*This lemma allows direct proofs of other <-properties*) +lemma zless_imp_succ_zadd_lemma: + "[| w $< z; w \ int; z \ int |] ==> (\n\nat. z = w $+ $#(succ(n)))" +apply (simp add: zless_def znegative_def zdiff_def int_def) +apply (auto dest!: less_imp_succ_add simp add: zadd zminus int_of_def) +apply (rule_tac x = k in bexI) +apply (erule_tac i="succ (v)" for v in add_left_cancel, auto) +done + +lemma zless_imp_succ_zadd: + "w $< z ==> (\n\nat. w $+ $#(succ(n)) = intify(z))" +apply (subgoal_tac "intify (w) $< intify (z) ") +apply (drule_tac w = "intify (w) " in zless_imp_succ_zadd_lemma) +apply auto +done + +lemma zless_succ_zadd_lemma: + "w \ int ==> w $< w $+ $# succ(n)" +apply (simp add: zless_def znegative_def zdiff_def int_def) +apply (auto simp add: zadd zminus int_of_def image_iff) +apply (rule_tac x = 0 in exI, auto) +done + +lemma zless_succ_zadd: "w $< w $+ $# succ(n)" +by (cut_tac intify_in_int [THEN zless_succ_zadd_lemma], auto) + +lemma zless_iff_succ_zadd: + "w $< z \ (\n\nat. w $+ $#(succ(n)) = intify(z))" +apply (rule iffI) +apply (erule zless_imp_succ_zadd, auto) +apply (rename_tac "n") +apply (cut_tac w = w and n = n in zless_succ_zadd, auto) +done + +lemma zless_int_of [simp]: "[| m\nat; n\nat |] ==> ($#m $< $#n) \ (m int; y \ int; z \ int |] ==> x $< z" +apply (simp add: zless_def znegative_def zdiff_def int_def) +apply (auto simp add: zadd zminus image_iff) +apply (rename_tac x1 x2 y1 y2) +apply (rule_tac x = "x1#+x2" in exI) +apply (rule_tac x = "y1#+y2" in exI) +apply (auto simp add: add_lt_mono) +apply (rule sym) +apply hypsubst_thin +apply (erule add_left_cancel)+ +apply auto +done + +lemma zless_trans [trans]: "[| x $< y; y $< z |] ==> x $< z" +apply (subgoal_tac "intify (x) $< intify (z) ") +apply (rule_tac [2] y = "intify (y) " in zless_trans_lemma) +apply auto +done + +lemma zless_not_sym: "z $< w ==> ~ (w $< z)" +by (blast dest: zless_trans) + +(* [| z $< w; ~ P ==> w $< z |] ==> P *) +lemmas zless_asym = zless_not_sym [THEN swap] + +lemma zless_imp_zle: "z $< w ==> z $\ w" +by (simp add: zle_def) + +lemma zle_linear: "z $\ w | w $\ z" +apply (simp add: zle_def) +apply (cut_tac zless_linear, blast) +done + + +subsection\Less Than or Equals\ + +lemma zle_refl: "z $\ z" +by (simp add: zle_def) + +lemma zle_eq_refl: "x=y ==> x $\ y" +by (simp add: zle_refl) + +lemma zle_anti_sym_intify: "[| x $\ y; y $\ x |] ==> intify(x) = intify(y)" +apply (simp add: zle_def, auto) +apply (blast dest: zless_trans) +done + +lemma zle_anti_sym: "[| x $\ y; y $\ x; x \ int; y \ int |] ==> x=y" +by (drule zle_anti_sym_intify, auto) + +lemma zle_trans_lemma: + "[| x \ int; y \ int; z \ int; x $\ y; y $\ z |] ==> x $\ z" +apply (simp add: zle_def, auto) +apply (blast intro: zless_trans) +done + +lemma zle_trans [trans]: "[| x $\ y; y $\ z |] ==> x $\ z" +apply (subgoal_tac "intify (x) $\ intify (z) ") +apply (rule_tac [2] y = "intify (y) " in zle_trans_lemma) +apply auto +done + +lemma zle_zless_trans [trans]: "[| i $\ j; j $< k |] ==> i $< k" +apply (auto simp add: zle_def) +apply (blast intro: zless_trans) +apply (simp add: zless_def zdiff_def zadd_def) +done + +lemma zless_zle_trans [trans]: "[| i $< j; j $\ k |] ==> i $< k" +apply (auto simp add: zle_def) +apply (blast intro: zless_trans) +apply (simp add: zless_def zdiff_def zminus_def) +done + +lemma not_zless_iff_zle: "~ (z $< w) \ (w $\ z)" +apply (cut_tac z = z and w = w in zless_linear) +apply (auto dest: zless_trans simp add: zle_def) +apply (auto dest!: zless_imp_intify_neq) +done + +lemma not_zle_iff_zless: "~ (z $\ w) \ (w $< z)" +by (simp add: not_zless_iff_zle [THEN iff_sym]) + + +subsection\More subtraction laws (for \zcompare_rls\)\ + +lemma zdiff_zdiff_eq: "(x $- y) $- z = x $- (y $+ z)" +by (simp add: zdiff_def zadd_ac) + +lemma zdiff_zdiff_eq2: "x $- (y $- z) = (x $+ z) $- y" +by (simp add: zdiff_def zadd_ac) + +lemma zdiff_zless_iff: "(x$-y $< z) \ (x $< z $+ y)" +by (simp add: zless_def zdiff_def zadd_ac) + +lemma zless_zdiff_iff: "(x $< z$-y) \ (x $+ y $< z)" +by (simp add: zless_def zdiff_def zadd_ac) + +lemma zdiff_eq_iff: "[| x \ int; z \ int |] ==> (x$-y = z) \ (x = z $+ y)" +by (auto simp add: zdiff_def zadd_assoc) + +lemma eq_zdiff_iff: "[| x \ int; z \ int |] ==> (x = z$-y) \ (x $+ y = z)" +by (auto simp add: zdiff_def zadd_assoc) + +lemma zdiff_zle_iff_lemma: + "[| x \ int; z \ int |] ==> (x$-y $\ z) \ (x $\ z $+ y)" +by (auto simp add: zle_def zdiff_eq_iff zdiff_zless_iff) + +lemma zdiff_zle_iff: "(x$-y $\ z) \ (x $\ z $+ y)" +by (cut_tac zdiff_zle_iff_lemma [OF intify_in_int intify_in_int], simp) + +lemma zle_zdiff_iff_lemma: + "[| x \ int; z \ int |] ==>(x $\ z$-y) \ (x $+ y $\ z)" +apply (auto simp add: zle_def zdiff_eq_iff zless_zdiff_iff) +apply (auto simp add: zdiff_def zadd_assoc) +done + +lemma zle_zdiff_iff: "(x $\ z$-y) \ (x $+ y $\ z)" +by (cut_tac zle_zdiff_iff_lemma [ OF intify_in_int intify_in_int], simp) + +text\This list of rewrites simplifies (in)equalities by bringing subtractions + to the top and then moving negative terms to the other side. + Use with \zadd_ac\\ +lemmas zcompare_rls = + zdiff_def [symmetric] + zadd_zdiff_eq zdiff_zadd_eq zdiff_zdiff_eq zdiff_zdiff_eq2 + zdiff_zless_iff zless_zdiff_iff zdiff_zle_iff zle_zdiff_iff + zdiff_eq_iff eq_zdiff_iff + + +subsection\Monotonicity and Cancellation Results for Instantiation + of the CancelNumerals Simprocs\ + +lemma zadd_left_cancel: + "[| w \ int; w': int |] ==> (z $+ w' = z $+ w) \ (w' = w)" +apply safe +apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) +apply (simp add: zadd_ac) +done + +lemma zadd_left_cancel_intify [simp]: + "(z $+ w' = z $+ w) \ intify(w') = intify(w)" +apply (rule iff_trans) +apply (rule_tac [2] zadd_left_cancel, auto) +done + +lemma zadd_right_cancel: + "[| w \ int; w': int |] ==> (w' $+ z = w $+ z) \ (w' = w)" +apply safe +apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) +apply (simp add: zadd_ac) +done + +lemma zadd_right_cancel_intify [simp]: + "(w' $+ z = w $+ z) \ intify(w') = intify(w)" +apply (rule iff_trans) +apply (rule_tac [2] zadd_right_cancel, auto) +done + +lemma zadd_right_cancel_zless [simp]: "(w' $+ z $< w $+ z) \ (w' $< w)" +by (simp add: zdiff_zless_iff [THEN iff_sym] zdiff_def zadd_assoc) + +lemma zadd_left_cancel_zless [simp]: "(z $+ w' $< z $+ w) \ (w' $< w)" +by (simp add: zadd_commute [of z] zadd_right_cancel_zless) + +lemma zadd_right_cancel_zle [simp]: "(w' $+ z $\ w $+ z) \ w' $\ w" +by (simp add: zle_def) + +lemma zadd_left_cancel_zle [simp]: "(z $+ w' $\ z $+ w) \ w' $\ w" +by (simp add: zadd_commute [of z] zadd_right_cancel_zle) + + +(*"v $\ w ==> v$+z $\ w$+z"*) +lemmas zadd_zless_mono1 = zadd_right_cancel_zless [THEN iffD2] + +(*"v $\ w ==> z$+v $\ z$+w"*) +lemmas zadd_zless_mono2 = zadd_left_cancel_zless [THEN iffD2] + +(*"v $\ w ==> v$+z $\ w$+z"*) +lemmas zadd_zle_mono1 = zadd_right_cancel_zle [THEN iffD2] + +(*"v $\ w ==> z$+v $\ z$+w"*) +lemmas zadd_zle_mono2 = zadd_left_cancel_zle [THEN iffD2] + +lemma zadd_zle_mono: "[| w' $\ w; z' $\ z |] ==> w' $+ z' $\ w $+ z" +by (erule zadd_zle_mono1 [THEN zle_trans], simp) + +lemma zadd_zless_mono: "[| w' $< w; z' $\ z |] ==> w' $+ z' $< w $+ z" +by (erule zadd_zless_mono1 [THEN zless_zle_trans], simp) + + +subsection\Comparison laws\ + +lemma zminus_zless_zminus [simp]: "($- x $< $- y) \ (y $< x)" +by (simp add: zless_def zdiff_def zadd_ac) + +lemma zminus_zle_zminus [simp]: "($- x $\ $- y) \ (y $\ x)" +by (simp add: not_zless_iff_zle [THEN iff_sym]) + +subsubsection\More inequality lemmas\ + +lemma equation_zminus: "[| x \ int; y \ int |] ==> (x = $- y) \ (y = $- x)" +by auto + +lemma zminus_equation: "[| x \ int; y \ int |] ==> ($- x = y) \ ($- y = x)" +by auto + +lemma equation_zminus_intify: "(intify(x) = $- y) \ (intify(y) = $- x)" +apply (cut_tac x = "intify (x) " and y = "intify (y) " in equation_zminus) +apply auto +done + +lemma zminus_equation_intify: "($- x = intify(y)) \ ($- y = intify(x))" +apply (cut_tac x = "intify (x) " and y = "intify (y) " in zminus_equation) +apply auto +done + + +subsubsection\The next several equations are permutative: watch out!\ + +lemma zless_zminus: "(x $< $- y) \ (y $< $- x)" +by (simp add: zless_def zdiff_def zadd_ac) + +lemma zminus_zless: "($- x $< y) \ ($- y $< x)" +by (simp add: zless_def zdiff_def zadd_ac) + +lemma zle_zminus: "(x $\ $- y) \ (y $\ $- x)" +by (simp add: not_zless_iff_zle [THEN iff_sym] zminus_zless) + +lemma zminus_zle: "($- x $\ y) \ ($- y $\ x)" +by (simp add: not_zless_iff_zle [THEN iff_sym] zless_zminus) + +end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/IntDiv.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/IntDiv.thy Sun Jun 24 15:57:48 2018 +0200 @@ -0,0 +1,1771 @@ +(* Title: ZF/IntDiv.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1999 University of Cambridge + +Here is the division algorithm in ML: + + fun posDivAlg (a,b) = + if a0 then posDivAlg (a,b) + else if a=0 then (0,0) + else negateSnd (negDivAlg (~a,~b)) + else + if 0The Division Operators Div and Mod\ + +theory IntDiv +imports Bin OrderArith +begin + +definition + quorem :: "[i,i] => o" where + "quorem == % . + a = b$*q $+ r & + (#0$r & r$ #0)" + +definition + adjust :: "[i,i] => i" where + "adjust(b) == %. if #0 $\ r$-b then <#2$*q $+ #1,r$-b> + else <#2$*q,r>" + + +(** the division algorithm **) + +definition + posDivAlg :: "i => i" where +(*for the case a>=0, b>0*) +(*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*) + "posDivAlg(ab) == + wfrec(measure(int*int, %. nat_of (a $- b $+ #1)), + ab, + % f. if (a$#0) then <#0,a> + else adjust(b, f ` ))" + + +(*for the case a<0, b>0*) +definition + negDivAlg :: "i => i" where +(*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*) + "negDivAlg(ab) == + wfrec(measure(int*int, %. nat_of ($- a $- b)), + ab, + % f. if (#0 $\ a$+b | b$\#0) then <#-1,a$+b> + else adjust(b, f ` ))" + +(*for the general case @{term"b\0"}*) + +definition + negateSnd :: "i => i" where + "negateSnd == %. " + + (*The full division algorithm considers all possible signs for a, b + including the special case a=0, b<0, because negDivAlg requires a<0*) +definition + divAlg :: "i => i" where + "divAlg == + %. if #0 $\ a then + if #0 $\ b then posDivAlg () + else if a=#0 then <#0,#0> + else negateSnd (negDivAlg (<$-a,$-b>)) + else + if #0$) + else negateSnd (posDivAlg (<$-a,$-b>))" + +definition + zdiv :: "[i,i]=>i" (infixl "zdiv" 70) where + "a zdiv b == fst (divAlg ())" + +definition + zmod :: "[i,i]=>i" (infixl "zmod" 70) where + "a zmod b == snd (divAlg ())" + + +(** Some basic laws by Sidi Ehmety (need linear arithmetic!) **) + +lemma zspos_add_zspos_imp_zspos: "[| #0 $< x; #0 $< y |] ==> #0 $< x $+ y" +apply (rule_tac y = "y" in zless_trans) +apply (rule_tac [2] zdiff_zless_iff [THEN iffD1]) +apply auto +done + +lemma zpos_add_zpos_imp_zpos: "[| #0 $\ x; #0 $\ y |] ==> #0 $\ x $+ y" +apply (rule_tac y = "y" in zle_trans) +apply (rule_tac [2] zdiff_zle_iff [THEN iffD1]) +apply auto +done + +lemma zneg_add_zneg_imp_zneg: "[| x $< #0; y $< #0 |] ==> x $+ y $< #0" +apply (rule_tac y = "y" in zless_trans) +apply (rule zless_zdiff_iff [THEN iffD1]) +apply auto +done + +(* this theorem is used below *) +lemma zneg_or_0_add_zneg_or_0_imp_zneg_or_0: + "[| x $\ #0; y $\ #0 |] ==> x $+ y $\ #0" +apply (rule_tac y = "y" in zle_trans) +apply (rule zle_zdiff_iff [THEN iffD1]) +apply auto +done + +lemma zero_lt_zmagnitude: "[| #0 $< k; k \ int |] ==> 0 < zmagnitude(k)" +apply (drule zero_zless_imp_znegative_zminus) +apply (drule_tac [2] zneg_int_of) +apply (auto simp add: zminus_equation [of k]) +apply (subgoal_tac "0 < zmagnitude ($# succ (n))") + apply simp +apply (simp only: zmagnitude_int_of) +apply simp +done + + +(*** Inequality lemmas involving $#succ(m) ***) + +lemma zless_add_succ_iff: + "(w $< z $+ $# succ(m)) \ (w $< z $+ $#m | intify(w) = z $+ $#m)" +apply (auto simp add: zless_iff_succ_zadd zadd_assoc int_of_add [symmetric]) +apply (rule_tac [3] x = "0" in bexI) +apply (cut_tac m = "m" in int_succ_int_1) +apply (cut_tac m = "n" in int_succ_int_1) +apply simp +apply (erule natE) +apply auto +apply (rule_tac x = "succ (n) " in bexI) +apply auto +done + +lemma zadd_succ_lemma: + "z \ int ==> (w $+ $# succ(m) $\ z) \ (w $+ $#m $< z)" +apply (simp only: not_zless_iff_zle [THEN iff_sym] zless_add_succ_iff) +apply (auto intro: zle_anti_sym elim: zless_asym + simp add: zless_imp_zle not_zless_iff_zle) +done + +lemma zadd_succ_zle_iff: "(w $+ $# succ(m) $\ z) \ (w $+ $#m $< z)" +apply (cut_tac z = "intify (z)" in zadd_succ_lemma) +apply auto +done + +(** Inequality reasoning **) + +lemma zless_add1_iff_zle: "(w $< z $+ #1) \ (w$\z)" +apply (subgoal_tac "#1 = $# 1") +apply (simp only: zless_add_succ_iff zle_def) +apply auto +done + +lemma add1_zle_iff: "(w $+ #1 $\ z) \ (w $< z)" +apply (subgoal_tac "#1 = $# 1") +apply (simp only: zadd_succ_zle_iff) +apply auto +done + +lemma add1_left_zle_iff: "(#1 $+ w $\ z) \ (w $< z)" +apply (subst zadd_commute) +apply (rule add1_zle_iff) +done + + +(*** Monotonicity of Multiplication ***) + +lemma zmult_mono_lemma: "k \ nat ==> i $\ j ==> i $* $#k $\ j $* $#k" +apply (induct_tac "k") + prefer 2 apply (subst int_succ_int_1) +apply (simp_all (no_asm_simp) add: zadd_zmult_distrib2 zadd_zle_mono) +done + +lemma zmult_zle_mono1: "[| i $\ j; #0 $\ k |] ==> i$*k $\ j$*k" +apply (subgoal_tac "i $* intify (k) $\ j $* intify (k) ") +apply (simp (no_asm_use)) +apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) +apply (rule_tac [3] zmult_mono_lemma) +apply auto +apply (simp add: znegative_iff_zless_0 not_zless_iff_zle [THEN iff_sym]) +done + +lemma zmult_zle_mono1_neg: "[| i $\ j; k $\ #0 |] ==> j$*k $\ i$*k" +apply (rule zminus_zle_zminus [THEN iffD1]) +apply (simp del: zmult_zminus_right + add: zmult_zminus_right [symmetric] zmult_zle_mono1 zle_zminus) +done + +lemma zmult_zle_mono2: "[| i $\ j; #0 $\ k |] ==> k$*i $\ k$*j" +apply (drule zmult_zle_mono1) +apply (simp_all add: zmult_commute) +done + +lemma zmult_zle_mono2_neg: "[| i $\ j; k $\ #0 |] ==> k$*j $\ k$*i" +apply (drule zmult_zle_mono1_neg) +apply (simp_all add: zmult_commute) +done + +(* $\ monotonicity, BOTH arguments*) +lemma zmult_zle_mono: + "[| i $\ j; k $\ l; #0 $\ j; #0 $\ k |] ==> i$*k $\ j$*l" +apply (erule zmult_zle_mono1 [THEN zle_trans]) +apply assumption +apply (erule zmult_zle_mono2) +apply assumption +done + + +(** strict, in 1st argument; proof is by induction on k>0 **) + +lemma zmult_zless_mono2_lemma [rule_format]: + "[| i$ nat |] ==> 0 $#k $* i $< $#k $* j" +apply (induct_tac "k") + prefer 2 + apply (subst int_succ_int_1) + apply (erule natE) +apply (simp_all add: zadd_zmult_distrib zadd_zless_mono zle_def) +apply (frule nat_0_le) +apply (subgoal_tac "i $+ (i $+ $# xa $* i) $< j $+ (j $+ $# xa $* j) ") +apply (simp (no_asm_use)) +apply (rule zadd_zless_mono) +apply (simp_all (no_asm_simp) add: zle_def) +done + +lemma zmult_zless_mono2: "[| i$ k$*i $< k$*j" +apply (subgoal_tac "intify (k) $* i $< intify (k) $* j") +apply (simp (no_asm_use)) +apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) +apply (rule_tac [3] zmult_zless_mono2_lemma) +apply auto +apply (simp add: znegative_iff_zless_0) +apply (drule zless_trans, assumption) +apply (auto simp add: zero_lt_zmagnitude) +done + +lemma zmult_zless_mono1: "[| i$ i$*k $< j$*k" +apply (drule zmult_zless_mono2) +apply (simp_all add: zmult_commute) +done + +(* < monotonicity, BOTH arguments*) +lemma zmult_zless_mono: + "[| i $< j; k $< l; #0 $< j; #0 $< k |] ==> i$*k $< j$*l" +apply (erule zmult_zless_mono1 [THEN zless_trans]) +apply assumption +apply (erule zmult_zless_mono2) +apply assumption +done + +lemma zmult_zless_mono1_neg: "[| i $< j; k $< #0 |] ==> j$*k $< i$*k" +apply (rule zminus_zless_zminus [THEN iffD1]) +apply (simp del: zmult_zminus_right + add: zmult_zminus_right [symmetric] zmult_zless_mono1 zless_zminus) +done + +lemma zmult_zless_mono2_neg: "[| i $< j; k $< #0 |] ==> k$*j $< k$*i" +apply (rule zminus_zless_zminus [THEN iffD1]) +apply (simp del: zmult_zminus + add: zmult_zminus [symmetric] zmult_zless_mono2 zless_zminus) +done + + +(** Products of zeroes **) + +lemma zmult_eq_lemma: + "[| m \ int; n \ int |] ==> (m = #0 | n = #0) \ (m$*n = #0)" +apply (case_tac "m $< #0") +apply (auto simp add: not_zless_iff_zle zle_def neq_iff_zless) +apply (force dest: zmult_zless_mono1_neg zmult_zless_mono1)+ +done + +lemma zmult_eq_0_iff [iff]: "(m$*n = #0) \ (intify(m) = #0 | intify(n) = #0)" +apply (simp add: zmult_eq_lemma) +done + + +(** Cancellation laws for k*m < k*n and m*k < n*k, also for @{text"\"} and =, + but not (yet?) for k*m < n*k. **) + +lemma zmult_zless_lemma: + "[| k \ int; m \ int; n \ int |] + ==> (m$*k $< n$*k) \ ((#0 $< k & m$ ((#0 $< k & m$ ((#0 $< k & m$ n$*k) \ ((#0 $< k \ m$\n) & (k $< #0 \ n$\m))" +by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel2) + +lemma zmult_zle_cancel1: + "(k$*m $\ k$*n) \ ((#0 $< k \ m$\n) & (k $< #0 \ n$\m))" +by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel1) + +lemma int_eq_iff_zle: "[| m \ int; n \ int |] ==> m=n \ (m $\ n & n $\ m)" +apply (blast intro: zle_refl zle_anti_sym) +done + +lemma zmult_cancel2_lemma: + "[| k \ int; m \ int; n \ int |] ==> (m$*k = n$*k) \ (k=#0 | m=n)" +apply (simp add: int_eq_iff_zle [of "m$*k"] int_eq_iff_zle [of m]) +apply (auto simp add: zmult_zle_cancel2 neq_iff_zless) +done + +lemma zmult_cancel2 [simp]: + "(m$*k = n$*k) \ (intify(k) = #0 | intify(m) = intify(n))" +apply (rule iff_trans) +apply (rule_tac [2] zmult_cancel2_lemma) +apply auto +done + +lemma zmult_cancel1 [simp]: + "(k$*m = k$*n) \ (intify(k) = #0 | intify(m) = intify(n))" +by (simp add: zmult_commute [of k] zmult_cancel2) + + +subsection\Uniqueness and monotonicity of quotients and remainders\ + +lemma unique_quotient_lemma: + "[| b$*q' $+ r' $\ b$*q $+ r; #0 $\ r'; #0 $< b; r $< b |] + ==> q' $\ q" +apply (subgoal_tac "r' $+ b $* (q'$-q) $\ r") + prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_ac zcompare_rls) +apply (subgoal_tac "#0 $< b $* (#1 $+ q $- q') ") + prefer 2 + apply (erule zle_zless_trans) + apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls) + apply (erule zle_zless_trans) + apply simp +apply (subgoal_tac "b $* q' $< b $* (#1 $+ q)") + prefer 2 + apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls) +apply (auto elim: zless_asym + simp add: zmult_zless_cancel1 zless_add1_iff_zle zadd_ac zcompare_rls) +done + +lemma unique_quotient_lemma_neg: + "[| b$*q' $+ r' $\ b$*q $+ r; r $\ #0; b $< #0; b $< r' |] + ==> q $\ q'" +apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r" + in unique_quotient_lemma) +apply (auto simp del: zminus_zadd_distrib + simp add: zminus_zadd_distrib [symmetric] zle_zminus zless_zminus) +done + + +lemma unique_quotient: + "[| quorem (, ); quorem (, ); b \ int; b \ #0; + q \ int; q' \ int |] ==> q = q'" +apply (simp add: split_ifs quorem_def neq_iff_zless) +apply safe +apply simp_all +apply (blast intro: zle_anti_sym + dest: zle_eq_refl [THEN unique_quotient_lemma] + zle_eq_refl [THEN unique_quotient_lemma_neg] sym)+ +done + +lemma unique_remainder: + "[| quorem (, ); quorem (, ); b \ int; b \ #0; + q \ int; q' \ int; + r \ int; r' \ int |] ==> r = r'" +apply (subgoal_tac "q = q'") + prefer 2 apply (blast intro: unique_quotient) +apply (simp add: quorem_def) +done + + +subsection\Correctness of posDivAlg, + the Division Algorithm for \a\0\ and \b>0\\ + +lemma adjust_eq [simp]: + "adjust(b, ) = (let diff = r$-b in + if #0 $\ diff then <#2$*q $+ #1,diff> + else <#2$*q,r>)" +by (simp add: Let_def adjust_def) + + +lemma posDivAlg_termination: + "[| #0 $< b; ~ a $< b |] + ==> nat_of(a $- #2 $* b $+ #1) < nat_of(a $- b $+ #1)" +apply (simp (no_asm) add: zless_nat_conj) +apply (simp add: not_zless_iff_zle zless_add1_iff_zle zcompare_rls) +done + +lemmas posDivAlg_unfold = def_wfrec [OF posDivAlg_def wf_measure] + +lemma posDivAlg_eqn: + "[| #0 $< b; a \ int; b \ int |] ==> + posDivAlg() = + (if a$ else adjust(b, posDivAlg ()))" +apply (rule posDivAlg_unfold [THEN trans]) +apply (simp add: vimage_iff not_zless_iff_zle [THEN iff_sym]) +apply (blast intro: posDivAlg_termination) +done + +lemma posDivAlg_induct_lemma [rule_format]: + assumes prem: + "!!a b. [| a \ int; b \ int; + ~ (a $< b | b $\ #0) \ P() |] ==> P()" + shows " \ int*int \ P()" +using wf_measure [where A = "int*int" and f = "%.nat_of (a $- b $+ #1)"] +proof (induct "" arbitrary: u v rule: wf_induct) + case (step x) + hence uv: "u \ int" "v \ int" by auto + thus ?case + apply (rule prem) + apply (rule impI) + apply (rule step) + apply (auto simp add: step uv not_zle_iff_zless posDivAlg_termination) + done +qed + + +lemma posDivAlg_induct [consumes 2]: + assumes u_int: "u \ int" + and v_int: "v \ int" + and ih: "!!a b. [| a \ int; b \ int; + ~ (a $< b | b $\ #0) \ P(a, #2 $* b) |] ==> P(a,b)" + shows "P(u,v)" +apply (subgoal_tac "(%. P (x,y)) ()") +apply simp +apply (rule posDivAlg_induct_lemma) +apply (simp (no_asm_use)) +apply (rule ih) +apply (auto simp add: u_int v_int) +done + +(*FIXME: use intify in integ_of so that we always have @{term"integ_of w \ int"}. + then this rewrite can work for all constants!!*) +lemma intify_eq_0_iff_zle: "intify(m) = #0 \ (m $\ #0 & #0 $\ m)" + by (simp add: int_eq_iff_zle) + + +subsection\Some convenient biconditionals for products of signs\ + +lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j" + by (drule zmult_zless_mono1, auto) + +lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j" + by (drule zmult_zless_mono1_neg, auto) + +lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0" + by (drule zmult_zless_mono1_neg, auto) + + +(** Inequality reasoning **) + +lemma int_0_less_lemma: + "[| x \ int; y \ int |] + ==> (#0 $< x $* y) \ (#0 $< x & #0 $< y | x $< #0 & y $< #0)" +apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg) +apply (rule ccontr) +apply (rule_tac [2] ccontr) +apply (auto simp add: zle_def not_zless_iff_zle) +apply (erule_tac P = "#0$< x$* y" in rev_mp) +apply (erule_tac [2] P = "#0$< x$* y" in rev_mp) +apply (drule zmult_pos_neg, assumption) + prefer 2 + apply (drule zmult_pos_neg, assumption) +apply (auto dest: zless_not_sym simp add: zmult_commute) +done + +lemma int_0_less_mult_iff: + "(#0 $< x $* y) \ (#0 $< x & #0 $< y | x $< #0 & y $< #0)" +apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_less_lemma) +apply auto +done + +lemma int_0_le_lemma: + "[| x \ int; y \ int |] + ==> (#0 $\ x $* y) \ (#0 $\ x & #0 $\ y | x $\ #0 & y $\ #0)" +by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff) + +lemma int_0_le_mult_iff: + "(#0 $\ x $* y) \ ((#0 $\ x & #0 $\ y) | (x $\ #0 & y $\ #0))" +apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_le_lemma) +apply auto +done + +lemma zmult_less_0_iff: + "(x $* y $< #0) \ (#0 $< x & y $< #0 | x $< #0 & #0 $< y)" +apply (auto simp add: int_0_le_mult_iff not_zle_iff_zless [THEN iff_sym]) +apply (auto dest: zless_not_sym simp add: not_zle_iff_zless) +done + +lemma zmult_le_0_iff: + "(x $* y $\ #0) \ (#0 $\ x & y $\ #0 | x $\ #0 & #0 $\ y)" +by (auto dest: zless_not_sym + simp add: int_0_less_mult_iff not_zless_iff_zle [THEN iff_sym]) + + +(*Typechecking for posDivAlg*) +lemma posDivAlg_type [rule_format]: + "[| a \ int; b \ int |] ==> posDivAlg() \ int * int" +apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) +apply assumption+ +apply (case_tac "#0 $< ba") + apply (simp add: posDivAlg_eqn adjust_def integ_of_type + split: split_if_asm) + apply clarify + apply (simp add: int_0_less_mult_iff not_zle_iff_zless) +apply (simp add: not_zless_iff_zle) +apply (subst posDivAlg_unfold) +apply simp +done + +(*Correctness of posDivAlg: it computes quotients correctly*) +lemma posDivAlg_correct [rule_format]: + "[| a \ int; b \ int |] + ==> #0 $\ a \ #0 $< b \ quorem (, posDivAlg())" +apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) +apply auto + apply (simp_all add: quorem_def) + txt\base case: a + apply (simp add: posDivAlg_eqn) + apply (simp add: not_zless_iff_zle [THEN iff_sym]) + apply (simp add: int_0_less_mult_iff) +txt\main argument\ +apply (subst posDivAlg_eqn) +apply (simp_all (no_asm_simp)) +apply (erule splitE) +apply (rule posDivAlg_type) +apply (simp_all add: int_0_less_mult_iff) +apply (auto simp add: zadd_zmult_distrib2 Let_def) +txt\now just linear arithmetic\ +apply (simp add: not_zle_iff_zless zdiff_zless_iff) +done + + +subsection\Correctness of negDivAlg, the division algorithm for a<0 and b>0\ + +lemma negDivAlg_termination: + "[| #0 $< b; a $+ b $< #0 |] + ==> nat_of($- a $- #2 $* b) < nat_of($- a $- b)" +apply (simp (no_asm) add: zless_nat_conj) +apply (simp add: zcompare_rls not_zle_iff_zless zless_zdiff_iff [THEN iff_sym] + zless_zminus) +done + +lemmas negDivAlg_unfold = def_wfrec [OF negDivAlg_def wf_measure] + +lemma negDivAlg_eqn: + "[| #0 $< b; a \ int; b \ int |] ==> + negDivAlg() = + (if #0 $\ a$+b then <#-1,a$+b> + else adjust(b, negDivAlg ()))" +apply (rule negDivAlg_unfold [THEN trans]) +apply (simp (no_asm_simp) add: vimage_iff not_zless_iff_zle [THEN iff_sym]) +apply (blast intro: negDivAlg_termination) +done + +lemma negDivAlg_induct_lemma [rule_format]: + assumes prem: + "!!a b. [| a \ int; b \ int; + ~ (#0 $\ a $+ b | b $\ #0) \ P() |] + ==> P()" + shows " \ int*int \ P()" +using wf_measure [where A = "int*int" and f = "%.nat_of ($- a $- b)"] +proof (induct "" arbitrary: u v rule: wf_induct) + case (step x) + hence uv: "u \ int" "v \ int" by auto + thus ?case + apply (rule prem) + apply (rule impI) + apply (rule step) + apply (auto simp add: step uv not_zle_iff_zless negDivAlg_termination) + done +qed + +lemma negDivAlg_induct [consumes 2]: + assumes u_int: "u \ int" + and v_int: "v \ int" + and ih: "!!a b. [| a \ int; b \ int; + ~ (#0 $\ a $+ b | b $\ #0) \ P(a, #2 $* b) |] + ==> P(a,b)" + shows "P(u,v)" +apply (subgoal_tac " (%. P (x,y)) ()") +apply simp +apply (rule negDivAlg_induct_lemma) +apply (simp (no_asm_use)) +apply (rule ih) +apply (auto simp add: u_int v_int) +done + + +(*Typechecking for negDivAlg*) +lemma negDivAlg_type: + "[| a \ int; b \ int |] ==> negDivAlg() \ int * int" +apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) +apply assumption+ +apply (case_tac "#0 $< ba") + apply (simp add: negDivAlg_eqn adjust_def integ_of_type + split: split_if_asm) + apply clarify + apply (simp add: int_0_less_mult_iff not_zle_iff_zless) +apply (simp add: not_zless_iff_zle) +apply (subst negDivAlg_unfold) +apply simp +done + + +(*Correctness of negDivAlg: it computes quotients correctly + It doesn't work if a=0 because the 0/b=0 rather than -1*) +lemma negDivAlg_correct [rule_format]: + "[| a \ int; b \ int |] + ==> a $< #0 \ #0 $< b \ quorem (, negDivAlg())" +apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) + apply auto + apply (simp_all add: quorem_def) + txt\base case: @{term "0$\a$+b"}\ + apply (simp add: negDivAlg_eqn) + apply (simp add: not_zless_iff_zle [THEN iff_sym]) + apply (simp add: int_0_less_mult_iff) +txt\main argument\ +apply (subst negDivAlg_eqn) +apply (simp_all (no_asm_simp)) +apply (erule splitE) +apply (rule negDivAlg_type) +apply (simp_all add: int_0_less_mult_iff) +apply (auto simp add: zadd_zmult_distrib2 Let_def) +txt\now just linear arithmetic\ +apply (simp add: not_zle_iff_zless zdiff_zless_iff) +done + + +subsection\Existence shown by proving the division algorithm to be correct\ + +(*the case a=0*) +lemma quorem_0: "[|b \ #0; b \ int|] ==> quorem (<#0,b>, <#0,#0>)" +by (force simp add: quorem_def neq_iff_zless) + +lemma posDivAlg_zero_divisor: "posDivAlg() = <#0,a>" +apply (subst posDivAlg_unfold) +apply simp +done + +lemma posDivAlg_0 [simp]: "posDivAlg (<#0,b>) = <#0,#0>" +apply (subst posDivAlg_unfold) +apply (simp add: not_zle_iff_zless) +done + + +(*Needed below. Actually it's an equivalence.*) +lemma linear_arith_lemma: "~ (#0 $\ #-1 $+ b) ==> (b $\ #0)" +apply (simp add: not_zle_iff_zless) +apply (drule zminus_zless_zminus [THEN iffD2]) +apply (simp add: zadd_commute zless_add1_iff_zle zle_zminus) +done + +lemma negDivAlg_minus1 [simp]: "negDivAlg (<#-1,b>) = <#-1, b$-#1>" +apply (subst negDivAlg_unfold) +apply (simp add: linear_arith_lemma integ_of_type vimage_iff) +done + +lemma negateSnd_eq [simp]: "negateSnd () = " +apply (unfold negateSnd_def) +apply auto +done + +lemma negateSnd_type: "qr \ int * int ==> negateSnd (qr) \ int * int" +apply (unfold negateSnd_def) +apply auto +done + +lemma quorem_neg: + "[|quorem (<$-a,$-b>, qr); a \ int; b \ int; qr \ int * int|] + ==> quorem (, negateSnd(qr))" +apply clarify +apply (auto elim: zless_asym simp add: quorem_def zless_zminus) +txt\linear arithmetic from here on\ +apply (simp_all add: zminus_equation [of a] zminus_zless) +apply (cut_tac [2] z = "b" and w = "#0" in zless_linear) +apply (cut_tac [1] z = "b" and w = "#0" in zless_linear) +apply auto +apply (blast dest: zle_zless_trans)+ +done + +lemma divAlg_correct: + "[|b \ #0; a \ int; b \ int|] ==> quorem (, divAlg())" +apply (auto simp add: quorem_0 divAlg_def) +apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct + posDivAlg_type negDivAlg_type) +apply (auto simp add: quorem_def neq_iff_zless) +txt\linear arithmetic from here on\ +apply (auto simp add: zle_def) +done + +lemma divAlg_type: "[|a \ int; b \ int|] ==> divAlg() \ int * int" +apply (auto simp add: divAlg_def) +apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type) +done + + +(** intify cancellation **) + +lemma zdiv_intify1 [simp]: "intify(x) zdiv y = x zdiv y" + by (simp add: zdiv_def) + +lemma zdiv_intify2 [simp]: "x zdiv intify(y) = x zdiv y" + by (simp add: zdiv_def) + +lemma zdiv_type [iff,TC]: "z zdiv w \ int" +apply (unfold zdiv_def) +apply (blast intro: fst_type divAlg_type) +done + +lemma zmod_intify1 [simp]: "intify(x) zmod y = x zmod y" + by (simp add: zmod_def) + +lemma zmod_intify2 [simp]: "x zmod intify(y) = x zmod y" + by (simp add: zmod_def) + +lemma zmod_type [iff,TC]: "z zmod w \ int" +apply (unfold zmod_def) +apply (rule snd_type) +apply (blast intro: divAlg_type) +done + + +(** Arbitrary definitions for division by zero. Useful to simplify + certain equations **) + +lemma DIVISION_BY_ZERO_ZDIV: "a zdiv #0 = #0" + by (simp add: zdiv_def divAlg_def posDivAlg_zero_divisor) + +lemma DIVISION_BY_ZERO_ZMOD: "a zmod #0 = intify(a)" + by (simp add: zmod_def divAlg_def posDivAlg_zero_divisor) + + +(** Basic laws about division and remainder **) + +lemma raw_zmod_zdiv_equality: + "[| a \ int; b \ int |] ==> a = b $* (a zdiv b) $+ (a zmod b)" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (cut_tac a = "a" and b = "b" in divAlg_correct) +apply (auto simp add: quorem_def zdiv_def zmod_def split_def) +done + +lemma zmod_zdiv_equality: "intify(a) = b $* (a zdiv b) $+ (a zmod b)" +apply (rule trans) +apply (rule_tac b = "intify (b)" in raw_zmod_zdiv_equality) +apply auto +done + +lemma pos_mod: "#0 $< b ==> #0 $\ a zmod b & a zmod b $< b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) +apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) +apply (blast dest: zle_zless_trans)+ +done + +lemmas pos_mod_sign = pos_mod [THEN conjunct1] + and pos_mod_bound = pos_mod [THEN conjunct2] + +lemma neg_mod: "b $< #0 ==> a zmod b $\ #0 & b $< a zmod b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) +apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) +apply (blast dest: zle_zless_trans) +apply (blast dest: zless_trans)+ +done + +lemmas neg_mod_sign = neg_mod [THEN conjunct1] + and neg_mod_bound = neg_mod [THEN conjunct2] + + +(** proving general properties of zdiv and zmod **) + +lemma quorem_div_mod: + "[|b \ #0; a \ int; b \ int |] + ==> quorem (, )" +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) +apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound + neg_mod_sign neg_mod_bound) +done + +(*Surely quorem(,) implies @{term"a \ int"}, but it doesn't matter*) +lemma quorem_div: + "[| quorem(,); b \ #0; a \ int; b \ int; q \ int |] + ==> a zdiv b = q" +by (blast intro: quorem_div_mod [THEN unique_quotient]) + +lemma quorem_mod: + "[| quorem(,); b \ #0; a \ int; b \ int; q \ int; r \ int |] + ==> a zmod b = r" +by (blast intro: quorem_div_mod [THEN unique_remainder]) + +lemma zdiv_pos_pos_trivial_raw: + "[| a \ int; b \ int; #0 $\ a; a $< b |] ==> a zdiv b = #0" +apply (rule quorem_div) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zle_zless_trans)+ +done + +lemma zdiv_pos_pos_trivial: "[| #0 $\ a; a $< b |] ==> a zdiv b = #0" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zdiv_pos_pos_trivial_raw) +apply auto +done + +lemma zdiv_neg_neg_trivial_raw: + "[| a \ int; b \ int; a $\ #0; b $< a |] ==> a zdiv b = #0" +apply (rule_tac r = "a" in quorem_div) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zle_zless_trans zless_trans)+ +done + +lemma zdiv_neg_neg_trivial: "[| a $\ #0; b $< a |] ==> a zdiv b = #0" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zdiv_neg_neg_trivial_raw) +apply auto +done + +lemma zadd_le_0_lemma: "[| a$+b $\ #0; #0 $< a; #0 $< b |] ==> False" +apply (drule_tac z' = "#0" and z = "b" in zadd_zless_mono) +apply (auto simp add: zle_def) +apply (blast dest: zless_trans) +done + +lemma zdiv_pos_neg_trivial_raw: + "[| a \ int; b \ int; #0 $< a; a$+b $\ #0 |] ==> a zdiv b = #-1" +apply (rule_tac r = "a $+ b" in quorem_div) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ +done + +lemma zdiv_pos_neg_trivial: "[| #0 $< a; a$+b $\ #0 |] ==> a zdiv b = #-1" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zdiv_pos_neg_trivial_raw) +apply auto +done + +(*There is no zdiv_neg_pos_trivial because #0 zdiv b = #0 would supersede it*) + + +lemma zmod_pos_pos_trivial_raw: + "[| a \ int; b \ int; #0 $\ a; a $< b |] ==> a zmod b = a" +apply (rule_tac q = "#0" in quorem_mod) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zle_zless_trans)+ +done + +lemma zmod_pos_pos_trivial: "[| #0 $\ a; a $< b |] ==> a zmod b = intify(a)" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zmod_pos_pos_trivial_raw) +apply auto +done + +lemma zmod_neg_neg_trivial_raw: + "[| a \ int; b \ int; a $\ #0; b $< a |] ==> a zmod b = a" +apply (rule_tac q = "#0" in quorem_mod) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zle_zless_trans zless_trans)+ +done + +lemma zmod_neg_neg_trivial: "[| a $\ #0; b $< a |] ==> a zmod b = intify(a)" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zmod_neg_neg_trivial_raw) +apply auto +done + +lemma zmod_pos_neg_trivial_raw: + "[| a \ int; b \ int; #0 $< a; a$+b $\ #0 |] ==> a zmod b = a$+b" +apply (rule_tac q = "#-1" in quorem_mod) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ +done + +lemma zmod_pos_neg_trivial: "[| #0 $< a; a$+b $\ #0 |] ==> a zmod b = a$+b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zmod_pos_neg_trivial_raw) +apply auto +done + +(*There is no zmod_neg_pos_trivial...*) + + +(*Simpler laws such as -a zdiv b = -(a zdiv b) FAIL*) + +lemma zdiv_zminus_zminus_raw: + "[|a \ int; b \ int|] ==> ($-a) zdiv ($-b) = a zdiv b" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_div]) +apply auto +done + +lemma zdiv_zminus_zminus [simp]: "($-a) zdiv ($-b) = a zdiv b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zminus_zminus_raw) +apply auto +done + +(*Simpler laws such as -a zmod b = -(a zmod b) FAIL*) +lemma zmod_zminus_zminus_raw: + "[|a \ int; b \ int|] ==> ($-a) zmod ($-b) = $- (a zmod b)" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod]) +apply auto +done + +lemma zmod_zminus_zminus [simp]: "($-a) zmod ($-b) = $- (a zmod b)" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zminus_zminus_raw) +apply auto +done + + +subsection\division of a number by itself\ + +lemma self_quotient_aux1: "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $\ q" +apply (subgoal_tac "#0 $< a$*q") +apply (cut_tac w = "#0" and z = "q" in add1_zle_iff) +apply (simp add: int_0_less_mult_iff) +apply (blast dest: zless_trans) +(*linear arithmetic...*) +apply (drule_tac t = "%x. x $- r" in subst_context) +apply (drule sym) +apply (simp add: zcompare_rls) +done + +lemma self_quotient_aux2: "[| #0 $< a; a = r $+ a$*q; #0 $\ r |] ==> q $\ #1" +apply (subgoal_tac "#0 $\ a$* (#1$-q)") + apply (simp add: int_0_le_mult_iff zcompare_rls) + apply (blast dest: zle_zless_trans) +apply (simp add: zdiff_zmult_distrib2) +apply (drule_tac t = "%x. x $- a $* q" in subst_context) +apply (simp add: zcompare_rls) +done + +lemma self_quotient: + "[| quorem(,); a \ int; q \ int; a \ #0|] ==> q = #1" +apply (simp add: split_ifs quorem_def neq_iff_zless) +apply (rule zle_anti_sym) +apply safe +apply auto +prefer 4 apply (blast dest: zless_trans) +apply (blast dest: zless_trans) +apply (rule_tac [3] a = "$-a" and r = "$-r" in self_quotient_aux1) +apply (rule_tac a = "$-a" and r = "$-r" in self_quotient_aux2) +apply (rule_tac [6] zminus_equation [THEN iffD1]) +apply (rule_tac [2] zminus_equation [THEN iffD1]) +apply (force intro: self_quotient_aux1 self_quotient_aux2 + simp add: zadd_commute zmult_zminus)+ +done + +lemma self_remainder: + "[|quorem(,); a \ int; q \ int; r \ int; a \ #0|] ==> r = #0" +apply (frule self_quotient) +apply (auto simp add: quorem_def) +done + +lemma zdiv_self_raw: "[|a \ #0; a \ int|] ==> a zdiv a = #1" +apply (blast intro: quorem_div_mod [THEN self_quotient]) +done + +lemma zdiv_self [simp]: "intify(a) \ #0 ==> a zdiv a = #1" +apply (drule zdiv_self_raw) +apply auto +done + +(*Here we have 0 zmod 0 = 0, also assumed by Knuth (who puts m zmod 0 = 0) *) +lemma zmod_self_raw: "a \ int ==> a zmod a = #0" +apply (case_tac "a = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (blast intro: quorem_div_mod [THEN self_remainder]) +done + +lemma zmod_self [simp]: "a zmod a = #0" +apply (cut_tac a = "intify (a)" in zmod_self_raw) +apply auto +done + + +subsection\Computation of division and remainder\ + +lemma zdiv_zero [simp]: "#0 zdiv b = #0" + by (simp add: zdiv_def divAlg_def) + +lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1" + by (simp (no_asm_simp) add: zdiv_def divAlg_def) + +lemma zmod_zero [simp]: "#0 zmod b = #0" + by (simp add: zmod_def divAlg_def) + +lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1" + by (simp add: zdiv_def divAlg_def) + +lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1" + by (simp add: zmod_def divAlg_def) + +(** a positive, b positive **) + +lemma zdiv_pos_pos: "[| #0 $< a; #0 $\ b |] + ==> a zdiv b = fst (posDivAlg())" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def) +apply (auto simp add: zle_def) +done + +lemma zmod_pos_pos: + "[| #0 $< a; #0 $\ b |] + ==> a zmod b = snd (posDivAlg())" +apply (simp (no_asm_simp) add: zmod_def divAlg_def) +apply (auto simp add: zle_def) +done + +(** a negative, b positive **) + +lemma zdiv_neg_pos: + "[| a $< #0; #0 $< b |] + ==> a zdiv b = fst (negDivAlg())" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def) +apply (blast dest: zle_zless_trans) +done + +lemma zmod_neg_pos: + "[| a $< #0; #0 $< b |] + ==> a zmod b = snd (negDivAlg())" +apply (simp (no_asm_simp) add: zmod_def divAlg_def) +apply (blast dest: zle_zless_trans) +done + +(** a positive, b negative **) + +lemma zdiv_pos_neg: + "[| #0 $< a; b $< #0 |] + ==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def intify_eq_0_iff_zle) +apply auto +apply (blast dest: zle_zless_trans)+ +apply (blast dest: zless_trans) +apply (blast intro: zless_imp_zle) +done + +lemma zmod_pos_neg: + "[| #0 $< a; b $< #0 |] + ==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))" +apply (simp (no_asm_simp) add: zmod_def divAlg_def intify_eq_0_iff_zle) +apply auto +apply (blast dest: zle_zless_trans)+ +apply (blast dest: zless_trans) +apply (blast intro: zless_imp_zle) +done + +(** a negative, b negative **) + +lemma zdiv_neg_neg: + "[| a $< #0; b $\ #0 |] + ==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def) +apply auto +apply (blast dest!: zle_zless_trans)+ +done + +lemma zmod_neg_neg: + "[| a $< #0; b $\ #0 |] + ==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))" +apply (simp (no_asm_simp) add: zmod_def divAlg_def) +apply auto +apply (blast dest!: zle_zless_trans)+ +done + +declare zdiv_pos_pos [of "integ_of (v)" "integ_of (w)", simp] for v w +declare zdiv_neg_pos [of "integ_of (v)" "integ_of (w)", simp] for v w +declare zdiv_pos_neg [of "integ_of (v)" "integ_of (w)", simp] for v w +declare zdiv_neg_neg [of "integ_of (v)" "integ_of (w)", simp] for v w +declare zmod_pos_pos [of "integ_of (v)" "integ_of (w)", simp] for v w +declare zmod_neg_pos [of "integ_of (v)" "integ_of (w)", simp] for v w +declare zmod_pos_neg [of "integ_of (v)" "integ_of (w)", simp] for v w +declare zmod_neg_neg [of "integ_of (v)" "integ_of (w)", simp] for v w +declare posDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", simp] for v w +declare negDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", simp] for v w + + +(** Special-case simplification **) + +lemma zmod_1 [simp]: "a zmod #1 = #0" +apply (cut_tac a = "a" and b = "#1" in pos_mod_sign) +apply (cut_tac [2] a = "a" and b = "#1" in pos_mod_bound) +apply auto +(*arithmetic*) +apply (drule add1_zle_iff [THEN iffD2]) +apply (rule zle_anti_sym) +apply auto +done + +lemma zdiv_1 [simp]: "a zdiv #1 = intify(a)" +apply (cut_tac a = "a" and b = "#1" in zmod_zdiv_equality) +apply auto +done + +lemma zmod_minus1_right [simp]: "a zmod #-1 = #0" +apply (cut_tac a = "a" and b = "#-1" in neg_mod_sign) +apply (cut_tac [2] a = "a" and b = "#-1" in neg_mod_bound) +apply auto +(*arithmetic*) +apply (drule add1_zle_iff [THEN iffD2]) +apply (rule zle_anti_sym) +apply auto +done + +lemma zdiv_minus1_right_raw: "a \ int ==> a zdiv #-1 = $-a" +apply (cut_tac a = "a" and b = "#-1" in zmod_zdiv_equality) +apply auto +apply (rule equation_zminus [THEN iffD2]) +apply auto +done + +lemma zdiv_minus1_right: "a zdiv #-1 = $-a" +apply (cut_tac a = "intify (a)" in zdiv_minus1_right_raw) +apply auto +done +declare zdiv_minus1_right [simp] + + +subsection\Monotonicity in the first argument (divisor)\ + +lemma zdiv_mono1: "[| a $\ a'; #0 $< b |] ==> a zdiv b $\ a' zdiv b" +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) +apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) +apply (rule unique_quotient_lemma) +apply (erule subst) +apply (erule subst) +apply (simp_all (no_asm_simp) add: pos_mod_sign pos_mod_bound) +done + +lemma zdiv_mono1_neg: "[| a $\ a'; b $< #0 |] ==> a' zdiv b $\ a zdiv b" +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) +apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) +apply (rule unique_quotient_lemma_neg) +apply (erule subst) +apply (erule subst) +apply (simp_all (no_asm_simp) add: neg_mod_sign neg_mod_bound) +done + + +subsection\Monotonicity in the second argument (dividend)\ + +lemma q_pos_lemma: + "[| #0 $\ b'$*q' $+ r'; r' $< b'; #0 $< b' |] ==> #0 $\ q'" +apply (subgoal_tac "#0 $< b'$* (q' $+ #1)") + apply (simp add: int_0_less_mult_iff) + apply (blast dest: zless_trans intro: zless_add1_iff_zle [THEN iffD1]) +apply (simp add: zadd_zmult_distrib2) +apply (erule zle_zless_trans) +apply (erule zadd_zless_mono2) +done + +lemma zdiv_mono2_lemma: + "[| b$*q $+ r = b'$*q' $+ r'; #0 $\ b'$*q' $+ r'; + r' $< b'; #0 $\ r; #0 $< b'; b' $\ b |] + ==> q $\ q'" +apply (frule q_pos_lemma, assumption+) +apply (subgoal_tac "b$*q $< b$* (q' $+ #1)") + apply (simp add: zmult_zless_cancel1) + apply (force dest: zless_add1_iff_zle [THEN iffD1] zless_trans zless_zle_trans) +apply (subgoal_tac "b$*q = r' $- r $+ b'$*q'") + prefer 2 apply (simp add: zcompare_rls) +apply (simp (no_asm_simp) add: zadd_zmult_distrib2) +apply (subst zadd_commute [of "b $* q'"], rule zadd_zless_mono) + prefer 2 apply (blast intro: zmult_zle_mono1) +apply (subgoal_tac "r' $+ #0 $< b $+ r") + apply (simp add: zcompare_rls) +apply (rule zadd_zless_mono) + apply auto +apply (blast dest: zless_zle_trans) +done + + +lemma zdiv_mono2_raw: + "[| #0 $\ a; #0 $< b'; b' $\ b; a \ int |] + ==> a zdiv b $\ a zdiv b'" +apply (subgoal_tac "#0 $< b") + prefer 2 apply (blast dest: zless_zle_trans) +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) +apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality) +apply (rule zdiv_mono2_lemma) +apply (erule subst) +apply (erule subst) +apply (simp_all add: pos_mod_sign pos_mod_bound) +done + +lemma zdiv_mono2: + "[| #0 $\ a; #0 $< b'; b' $\ b |] + ==> a zdiv b $\ a zdiv b'" +apply (cut_tac a = "intify (a)" in zdiv_mono2_raw) +apply auto +done + +lemma q_neg_lemma: + "[| b'$*q' $+ r' $< #0; #0 $\ r'; #0 $< b' |] ==> q' $< #0" +apply (subgoal_tac "b'$*q' $< #0") + prefer 2 apply (force intro: zle_zless_trans) +apply (simp add: zmult_less_0_iff) +apply (blast dest: zless_trans) +done + + + +lemma zdiv_mono2_neg_lemma: + "[| b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0; + r $< b; #0 $\ r'; #0 $< b'; b' $\ b |] + ==> q' $\ q" +apply (subgoal_tac "#0 $< b") + prefer 2 apply (blast dest: zless_zle_trans) +apply (frule q_neg_lemma, assumption+) +apply (subgoal_tac "b$*q' $< b$* (q $+ #1)") + apply (simp add: zmult_zless_cancel1) + apply (blast dest: zless_trans zless_add1_iff_zle [THEN iffD1]) +apply (simp (no_asm_simp) add: zadd_zmult_distrib2) +apply (subgoal_tac "b$*q' $\ b'$*q'") + prefer 2 + apply (simp add: zmult_zle_cancel2) + apply (blast dest: zless_trans) +apply (subgoal_tac "b'$*q' $+ r $< b $+ (b$*q $+ r)") + prefer 2 + apply (erule ssubst) + apply simp + apply (drule_tac w' = "r" and z' = "#0" in zadd_zless_mono) + apply (assumption) + apply simp +apply (simp (no_asm_use) add: zadd_commute) +apply (rule zle_zless_trans) + prefer 2 apply (assumption) +apply (simp (no_asm_simp) add: zmult_zle_cancel2) +apply (blast dest: zless_trans) +done + +lemma zdiv_mono2_neg_raw: + "[| a $< #0; #0 $< b'; b' $\ b; a \ int |] + ==> a zdiv b' $\ a zdiv b" +apply (subgoal_tac "#0 $< b") + prefer 2 apply (blast dest: zless_zle_trans) +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) +apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality) +apply (rule zdiv_mono2_neg_lemma) +apply (erule subst) +apply (erule subst) +apply (simp_all add: pos_mod_sign pos_mod_bound) +done + +lemma zdiv_mono2_neg: "[| a $< #0; #0 $< b'; b' $\ b |] + ==> a zdiv b' $\ a zdiv b" +apply (cut_tac a = "intify (a)" in zdiv_mono2_neg_raw) +apply auto +done + + + +subsection\More algebraic laws for zdiv and zmod\ + +(** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **) + +lemma zmult1_lemma: + "[| quorem(, ); c \ int; c \ #0 |] + ==> quorem (, )" +apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 + pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) +apply (auto intro: raw_zmod_zdiv_equality) +done + +lemma zdiv_zmult1_eq_raw: + "[|b \ int; c \ int|] + ==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" +apply (case_tac "c = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_div]) +apply auto +done + +lemma zdiv_zmult1_eq: "(a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" +apply (cut_tac b = "intify (b)" and c = "intify (c)" in zdiv_zmult1_eq_raw) +apply auto +done + +lemma zmod_zmult1_eq_raw: + "[|b \ int; c \ int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c" +apply (case_tac "c = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod]) +apply auto +done + +lemma zmod_zmult1_eq: "(a$*b) zmod c = a$*(b zmod c) zmod c" +apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult1_eq_raw) +apply auto +done + +lemma zmod_zmult1_eq': "(a$*b) zmod c = ((a zmod c) $* b) zmod c" +apply (rule trans) +apply (rule_tac b = " (b $* a) zmod c" in trans) +apply (rule_tac [2] zmod_zmult1_eq) +apply (simp_all (no_asm) add: zmult_commute) +done + +lemma zmod_zmult_distrib: "(a$*b) zmod c = ((a zmod c) $* (b zmod c)) zmod c" +apply (rule zmod_zmult1_eq' [THEN trans]) +apply (rule zmod_zmult1_eq) +done + +lemma zdiv_zmult_self1 [simp]: "intify(b) \ #0 ==> (a$*b) zdiv b = intify(a)" + by (simp add: zdiv_zmult1_eq) + +lemma zdiv_zmult_self2 [simp]: "intify(b) \ #0 ==> (b$*a) zdiv b = intify(a)" + by (simp add: zmult_commute) + +lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0" + by (simp add: zmod_zmult1_eq) + +lemma zmod_zmult_self2 [simp]: "(b$*a) zmod b = #0" + by (simp add: zmult_commute zmod_zmult1_eq) + + +(** proving (a$+b) zdiv c = + a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **) + +lemma zadd1_lemma: + "[| quorem(, ); quorem(, ); + c \ int; c \ #0 |] + ==> quorem (, )" +apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 + pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) +apply (auto intro: raw_zmod_zdiv_equality) +done + +(*NOT suitable for rewriting: the RHS has an instance of the LHS*) +lemma zdiv_zadd1_eq_raw: + "[|a \ int; b \ int; c \ int|] ==> + (a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" +apply (case_tac "c = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, + THEN quorem_div]) +done + +lemma zdiv_zadd1_eq: + "(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" +apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" + in zdiv_zadd1_eq_raw) +apply auto +done + +lemma zmod_zadd1_eq_raw: + "[|a \ int; b \ int; c \ int|] + ==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" +apply (case_tac "c = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, + THEN quorem_mod]) +done + +lemma zmod_zadd1_eq: "(a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" +apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" + in zmod_zadd1_eq_raw) +apply auto +done + +lemma zmod_div_trivial_raw: + "[|a \ int; b \ int|] ==> (a zmod b) zdiv b = #0" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound + zdiv_pos_pos_trivial neg_mod_sign neg_mod_bound zdiv_neg_neg_trivial) +done + +lemma zmod_div_trivial [simp]: "(a zmod b) zdiv b = #0" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_div_trivial_raw) +apply auto +done + +lemma zmod_mod_trivial_raw: + "[|a \ int; b \ int|] ==> (a zmod b) zmod b = a zmod b" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound + zmod_pos_pos_trivial neg_mod_sign neg_mod_bound zmod_neg_neg_trivial) +done + +lemma zmod_mod_trivial [simp]: "(a zmod b) zmod b = a zmod b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_mod_trivial_raw) +apply auto +done + +lemma zmod_zadd_left_eq: "(a$+b) zmod c = ((a zmod c) $+ b) zmod c" +apply (rule trans [symmetric]) +apply (rule zmod_zadd1_eq) +apply (simp (no_asm)) +apply (rule zmod_zadd1_eq [symmetric]) +done + +lemma zmod_zadd_right_eq: "(a$+b) zmod c = (a $+ (b zmod c)) zmod c" +apply (rule trans [symmetric]) +apply (rule zmod_zadd1_eq) +apply (simp (no_asm)) +apply (rule zmod_zadd1_eq [symmetric]) +done + + +lemma zdiv_zadd_self1 [simp]: + "intify(a) \ #0 ==> (a$+b) zdiv a = b zdiv a $+ #1" +by (simp (no_asm_simp) add: zdiv_zadd1_eq) + +lemma zdiv_zadd_self2 [simp]: + "intify(a) \ #0 ==> (b$+a) zdiv a = b zdiv a $+ #1" +by (simp (no_asm_simp) add: zdiv_zadd1_eq) + +lemma zmod_zadd_self1 [simp]: "(a$+b) zmod a = b zmod a" +apply (case_tac "a = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (simp (no_asm_simp) add: zmod_zadd1_eq) +done + +lemma zmod_zadd_self2 [simp]: "(b$+a) zmod a = b zmod a" +apply (case_tac "a = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (simp (no_asm_simp) add: zmod_zadd1_eq) +done + + +subsection\proving a zdiv (b*c) = (a zdiv b) zdiv c\ + +(*The condition c>0 seems necessary. Consider that 7 zdiv ~6 = ~2 but + 7 zdiv 2 zdiv ~3 = 3 zdiv ~3 = ~1. The subcase (a zdiv b) zmod c = 0 seems + to cause particular problems.*) + +(** first, four lemmas to bound the remainder for the cases b<0 and b>0 **) + +lemma zdiv_zmult2_aux1: + "[| #0 $< c; b $< r; r $\ #0 |] ==> b$*c $< b$*(q zmod c) $+ r" +apply (subgoal_tac "b $* (c $- q zmod c) $< r $* #1") +apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) +apply (rule zle_zless_trans) +apply (erule_tac [2] zmult_zless_mono1) +apply (rule zmult_zle_mono2_neg) +apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound) +apply (blast intro: zless_imp_zle dest: zless_zle_trans) +done + +lemma zdiv_zmult2_aux2: + "[| #0 $< c; b $< r; r $\ #0 |] ==> b $* (q zmod c) $+ r $\ #0" +apply (subgoal_tac "b $* (q zmod c) $\ #0") + prefer 2 + apply (simp add: zmult_le_0_iff pos_mod_sign) + apply (blast intro: zless_imp_zle dest: zless_zle_trans) +(*arithmetic*) +apply (drule zadd_zle_mono) +apply assumption +apply (simp add: zadd_commute) +done + +lemma zdiv_zmult2_aux3: + "[| #0 $< c; #0 $\ r; r $< b |] ==> #0 $\ b $* (q zmod c) $+ r" +apply (subgoal_tac "#0 $\ b $* (q zmod c)") + prefer 2 + apply (simp add: int_0_le_mult_iff pos_mod_sign) + apply (blast intro: zless_imp_zle dest: zle_zless_trans) +(*arithmetic*) +apply (drule zadd_zle_mono) +apply assumption +apply (simp add: zadd_commute) +done + +lemma zdiv_zmult2_aux4: + "[| #0 $< c; #0 $\ r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c" +apply (subgoal_tac "r $* #1 $< b $* (c $- q zmod c)") +apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) +apply (rule zless_zle_trans) +apply (erule zmult_zless_mono1) +apply (rule_tac [2] zmult_zle_mono2) +apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound) +apply (blast intro: zless_imp_zle dest: zle_zless_trans) +done + +lemma zdiv_zmult2_lemma: + "[| quorem (, ); a \ int; b \ int; b \ #0; #0 $< c |] + ==> quorem (, )" +apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def + neq_iff_zless int_0_less_mult_iff + zadd_zmult_distrib2 [symmetric] zdiv_zmult2_aux1 zdiv_zmult2_aux2 + zdiv_zmult2_aux3 zdiv_zmult2_aux4) +apply (blast dest: zless_trans)+ +done + +lemma zdiv_zmult2_eq_raw: + "[|#0 $< c; a \ int; b \ int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_div]) +apply (auto simp add: intify_eq_0_iff_zle) +apply (blast dest: zle_zless_trans) +done + +lemma zdiv_zmult2_eq: "#0 $< c ==> a zdiv (b$*c) = (a zdiv b) zdiv c" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zmult2_eq_raw) +apply auto +done + +lemma zmod_zmult2_eq_raw: + "[|#0 $< c; a \ int; b \ int|] + ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_mod]) +apply (auto simp add: intify_eq_0_iff_zle) +apply (blast dest: zle_zless_trans) +done + +lemma zmod_zmult2_eq: + "#0 $< c ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zmult2_eq_raw) +apply auto +done + +subsection\Cancellation of common factors in "zdiv"\ + +lemma zdiv_zmult_zmult1_aux1: + "[| #0 $< b; intify(c) \ #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" +apply (subst zdiv_zmult2_eq) +apply auto +done + +lemma zdiv_zmult_zmult1_aux2: + "[| b $< #0; intify(c) \ #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" +apply (subgoal_tac " (c $* ($-a)) zdiv (c $* ($-b)) = ($-a) zdiv ($-b)") +apply (rule_tac [2] zdiv_zmult_zmult1_aux1) +apply auto +done + +lemma zdiv_zmult_zmult1_raw: + "[|intify(c) \ #0; b \ int|] ==> (c$*a) zdiv (c$*b) = a zdiv b" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (auto simp add: neq_iff_zless [of b] + zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2) +done + +lemma zdiv_zmult_zmult1: "intify(c) \ #0 ==> (c$*a) zdiv (c$*b) = a zdiv b" +apply (cut_tac b = "intify (b)" in zdiv_zmult_zmult1_raw) +apply auto +done + +lemma zdiv_zmult_zmult2: "intify(c) \ #0 ==> (a$*c) zdiv (b$*c) = a zdiv b" +apply (drule zdiv_zmult_zmult1) +apply (auto simp add: zmult_commute) +done + + +subsection\Distribution of factors over "zmod"\ + +lemma zmod_zmult_zmult1_aux1: + "[| #0 $< b; intify(c) \ #0 |] + ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" +apply (subst zmod_zmult2_eq) +apply auto +done + +lemma zmod_zmult_zmult1_aux2: + "[| b $< #0; intify(c) \ #0 |] + ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" +apply (subgoal_tac " (c $* ($-a)) zmod (c $* ($-b)) = c $* (($-a) zmod ($-b))") +apply (rule_tac [2] zmod_zmult_zmult1_aux1) +apply auto +done + +lemma zmod_zmult_zmult1_raw: + "[|b \ int; c \ int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (case_tac "c = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (auto simp add: neq_iff_zless [of b] + zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2) +done + +lemma zmod_zmult_zmult1: "(c$*a) zmod (c$*b) = c $* (a zmod b)" +apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult_zmult1_raw) +apply auto +done + +lemma zmod_zmult_zmult2: "(a$*c) zmod (b$*c) = (a zmod b) $* c" +apply (cut_tac c = "c" in zmod_zmult_zmult1) +apply (auto simp add: zmult_commute) +done + + +(** Quotients of signs **) + +lemma zdiv_neg_pos_less0: "[| a $< #0; #0 $< b |] ==> a zdiv b $< #0" +apply (subgoal_tac "a zdiv b $\ #-1") +apply (erule zle_zless_trans) +apply (simp (no_asm)) +apply (rule zle_trans) +apply (rule_tac a' = "#-1" in zdiv_mono1) +apply (rule zless_add1_iff_zle [THEN iffD1]) +apply (simp (no_asm)) +apply (auto simp add: zdiv_minus1) +done + +lemma zdiv_nonneg_neg_le0: "[| #0 $\ a; b $< #0 |] ==> a zdiv b $\ #0" +apply (drule zdiv_mono1_neg) +apply auto +done + +lemma pos_imp_zdiv_nonneg_iff: "#0 $< b ==> (#0 $\ a zdiv b) \ (#0 $\ a)" +apply auto +apply (drule_tac [2] zdiv_mono1) +apply (auto simp add: neq_iff_zless) +apply (simp (no_asm_use) add: not_zless_iff_zle [THEN iff_sym]) +apply (blast intro: zdiv_neg_pos_less0) +done + +lemma neg_imp_zdiv_nonneg_iff: "b $< #0 ==> (#0 $\ a zdiv b) \ (a $\ #0)" +apply (subst zdiv_zminus_zminus [symmetric]) +apply (rule iff_trans) +apply (rule pos_imp_zdiv_nonneg_iff) +apply auto +done + +(*But not (a zdiv b $\ 0 iff a$\0); consider a=1, b=2 when a zdiv b = 0.*) +lemma pos_imp_zdiv_neg_iff: "#0 $< b ==> (a zdiv b $< #0) \ (a $< #0)" +apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) +apply (erule pos_imp_zdiv_nonneg_iff) +done + +(*Again the law fails for $\: consider a = -1, b = -2 when a zdiv b = 0*) +lemma neg_imp_zdiv_neg_iff: "b $< #0 ==> (a zdiv b $< #0) \ (#0 $< a)" +apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) +apply (erule neg_imp_zdiv_nonneg_iff) +done + +(* + THESE REMAIN TO BE CONVERTED -- but aren't that useful! + + subsection{* Speeding up the division algorithm with shifting *} + + (** computing "zdiv" by shifting **) + + lemma pos_zdiv_mult_2: "#0 $\ a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a" + apply (case_tac "a = #0") + apply (subgoal_tac "#1 $\ a") + apply (arith_tac 2) + apply (subgoal_tac "#1 $< a $* #2") + apply (arith_tac 2) + apply (subgoal_tac "#2$* (#1 $+ b zmod a) $\ #2$*a") + apply (rule_tac [2] zmult_zle_mono2) + apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound) + apply (subst zdiv_zadd1_eq) + apply (simp (no_asm_simp) add: zdiv_zmult_zmult2 zmod_zmult_zmult2 zdiv_pos_pos_trivial) + apply (subst zdiv_pos_pos_trivial) + apply (simp (no_asm_simp) add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ]) + apply (auto simp add: zmod_pos_pos_trivial) + apply (subgoal_tac "#0 $\ b zmod a") + apply (asm_simp_tac (simpset () add: pos_mod_sign) 2) + apply arith + done + + + lemma neg_zdiv_mult_2: "a $\ #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) \ (b$+#1) zdiv a" + apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zdiv (#2 $* ($-a)) \ ($-b-#1) zdiv ($-a)") + apply (rule_tac [2] pos_zdiv_mult_2) + apply (auto simp add: zmult_zminus_right) + apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))") + apply (Simp_tac 2) + apply (asm_full_simp_tac (HOL_ss add: zdiv_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) + done + + + (*Not clear why this must be proved separately; probably integ_of causes + simplification problems*) + lemma lemma: "~ #0 $\ x ==> x $\ #0" + apply auto + done + + lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) = + (if ~b | #0 $\ integ_of w + then integ_of v zdiv (integ_of w) + else (integ_of v $+ #1) zdiv (integ_of w))" + apply (simp_tac (simpset_of @{theory_context Int} add: zadd_assoc integ_of_BIT) + apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zdiv_zmult_zmult1 pos_zdiv_mult_2 lemma neg_zdiv_mult_2) + done + + declare zdiv_integ_of_BIT [simp] + + + (** computing "zmod" by shifting (proofs resemble those for "zdiv") **) + + lemma pos_zmod_mult_2: "#0 $\ a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)" + apply (case_tac "a = #0") + apply (subgoal_tac "#1 $\ a") + apply (arith_tac 2) + apply (subgoal_tac "#1 $< a $* #2") + apply (arith_tac 2) + apply (subgoal_tac "#2$* (#1 $+ b zmod a) $\ #2$*a") + apply (rule_tac [2] zmult_zle_mono2) + apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound) + apply (subst zmod_zadd1_eq) + apply (simp (no_asm_simp) add: zmod_zmult_zmult2 zmod_pos_pos_trivial) + apply (rule zmod_pos_pos_trivial) + apply (simp (no_asm_simp) # add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ]) + apply (auto simp add: zmod_pos_pos_trivial) + apply (subgoal_tac "#0 $\ b zmod a") + apply (asm_simp_tac (simpset () add: pos_mod_sign) 2) + apply arith + done + + + lemma neg_zmod_mult_2: "a $\ #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1" + apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zmod (#2$* ($-a)) = #1 $+ #2$* (($-b-#1) zmod ($-a))") + apply (rule_tac [2] pos_zmod_mult_2) + apply (auto simp add: zmult_zminus_right) + apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))") + apply (Simp_tac 2) + apply (asm_full_simp_tac (HOL_ss add: zmod_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) + apply (dtac (zminus_equation [THEN iffD1, symmetric]) + apply auto + done + + lemma zmod_integ_of_BIT: "integ_of (v BIT b) zmod integ_of (w BIT False) = + (if b then + if #0 $\ integ_of w + then #2 $* (integ_of v zmod integ_of w) $+ #1 + else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1 + else #2 $* (integ_of v zmod integ_of w))" + apply (simp_tac (simpset_of @{theory_context Int} add: zadd_assoc integ_of_BIT) + apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zmod_zmult_zmult1 pos_zmod_mult_2 lemma neg_zmod_mult_2) + done + + declare zmod_integ_of_BIT [simp] +*) + +end + diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/IntDiv_ZF.thy --- a/src/ZF/IntDiv_ZF.thy Sun Jun 24 11:41:41 2018 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1771 +0,0 @@ -(* Title: ZF/IntDiv_ZF.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1999 University of Cambridge - -Here is the division algorithm in ML: - - fun posDivAlg (a,b) = - if a0 then posDivAlg (a,b) - else if a=0 then (0,0) - else negateSnd (negDivAlg (~a,~b)) - else - if 0The Division Operators Div and Mod\ - -theory IntDiv_ZF -imports Bin OrderArith -begin - -definition - quorem :: "[i,i] => o" where - "quorem == % . - a = b$*q $+ r & - (#0$r & r$ #0)" - -definition - adjust :: "[i,i] => i" where - "adjust(b) == %. if #0 $\ r$-b then <#2$*q $+ #1,r$-b> - else <#2$*q,r>" - - -(** the division algorithm **) - -definition - posDivAlg :: "i => i" where -(*for the case a>=0, b>0*) -(*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*) - "posDivAlg(ab) == - wfrec(measure(int*int, %. nat_of (a $- b $+ #1)), - ab, - % f. if (a$#0) then <#0,a> - else adjust(b, f ` ))" - - -(*for the case a<0, b>0*) -definition - negDivAlg :: "i => i" where -(*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*) - "negDivAlg(ab) == - wfrec(measure(int*int, %. nat_of ($- a $- b)), - ab, - % f. if (#0 $\ a$+b | b$\#0) then <#-1,a$+b> - else adjust(b, f ` ))" - -(*for the general case @{term"b\0"}*) - -definition - negateSnd :: "i => i" where - "negateSnd == %. " - - (*The full division algorithm considers all possible signs for a, b - including the special case a=0, b<0, because negDivAlg requires a<0*) -definition - divAlg :: "i => i" where - "divAlg == - %. if #0 $\ a then - if #0 $\ b then posDivAlg () - else if a=#0 then <#0,#0> - else negateSnd (negDivAlg (<$-a,$-b>)) - else - if #0$) - else negateSnd (posDivAlg (<$-a,$-b>))" - -definition - zdiv :: "[i,i]=>i" (infixl "zdiv" 70) where - "a zdiv b == fst (divAlg ())" - -definition - zmod :: "[i,i]=>i" (infixl "zmod" 70) where - "a zmod b == snd (divAlg ())" - - -(** Some basic laws by Sidi Ehmety (need linear arithmetic!) **) - -lemma zspos_add_zspos_imp_zspos: "[| #0 $< x; #0 $< y |] ==> #0 $< x $+ y" -apply (rule_tac y = "y" in zless_trans) -apply (rule_tac [2] zdiff_zless_iff [THEN iffD1]) -apply auto -done - -lemma zpos_add_zpos_imp_zpos: "[| #0 $\ x; #0 $\ y |] ==> #0 $\ x $+ y" -apply (rule_tac y = "y" in zle_trans) -apply (rule_tac [2] zdiff_zle_iff [THEN iffD1]) -apply auto -done - -lemma zneg_add_zneg_imp_zneg: "[| x $< #0; y $< #0 |] ==> x $+ y $< #0" -apply (rule_tac y = "y" in zless_trans) -apply (rule zless_zdiff_iff [THEN iffD1]) -apply auto -done - -(* this theorem is used below *) -lemma zneg_or_0_add_zneg_or_0_imp_zneg_or_0: - "[| x $\ #0; y $\ #0 |] ==> x $+ y $\ #0" -apply (rule_tac y = "y" in zle_trans) -apply (rule zle_zdiff_iff [THEN iffD1]) -apply auto -done - -lemma zero_lt_zmagnitude: "[| #0 $< k; k \ int |] ==> 0 < zmagnitude(k)" -apply (drule zero_zless_imp_znegative_zminus) -apply (drule_tac [2] zneg_int_of) -apply (auto simp add: zminus_equation [of k]) -apply (subgoal_tac "0 < zmagnitude ($# succ (n))") - apply simp -apply (simp only: zmagnitude_int_of) -apply simp -done - - -(*** Inequality lemmas involving $#succ(m) ***) - -lemma zless_add_succ_iff: - "(w $< z $+ $# succ(m)) \ (w $< z $+ $#m | intify(w) = z $+ $#m)" -apply (auto simp add: zless_iff_succ_zadd zadd_assoc int_of_add [symmetric]) -apply (rule_tac [3] x = "0" in bexI) -apply (cut_tac m = "m" in int_succ_int_1) -apply (cut_tac m = "n" in int_succ_int_1) -apply simp -apply (erule natE) -apply auto -apply (rule_tac x = "succ (n) " in bexI) -apply auto -done - -lemma zadd_succ_lemma: - "z \ int ==> (w $+ $# succ(m) $\ z) \ (w $+ $#m $< z)" -apply (simp only: not_zless_iff_zle [THEN iff_sym] zless_add_succ_iff) -apply (auto intro: zle_anti_sym elim: zless_asym - simp add: zless_imp_zle not_zless_iff_zle) -done - -lemma zadd_succ_zle_iff: "(w $+ $# succ(m) $\ z) \ (w $+ $#m $< z)" -apply (cut_tac z = "intify (z)" in zadd_succ_lemma) -apply auto -done - -(** Inequality reasoning **) - -lemma zless_add1_iff_zle: "(w $< z $+ #1) \ (w$\z)" -apply (subgoal_tac "#1 = $# 1") -apply (simp only: zless_add_succ_iff zle_def) -apply auto -done - -lemma add1_zle_iff: "(w $+ #1 $\ z) \ (w $< z)" -apply (subgoal_tac "#1 = $# 1") -apply (simp only: zadd_succ_zle_iff) -apply auto -done - -lemma add1_left_zle_iff: "(#1 $+ w $\ z) \ (w $< z)" -apply (subst zadd_commute) -apply (rule add1_zle_iff) -done - - -(*** Monotonicity of Multiplication ***) - -lemma zmult_mono_lemma: "k \ nat ==> i $\ j ==> i $* $#k $\ j $* $#k" -apply (induct_tac "k") - prefer 2 apply (subst int_succ_int_1) -apply (simp_all (no_asm_simp) add: zadd_zmult_distrib2 zadd_zle_mono) -done - -lemma zmult_zle_mono1: "[| i $\ j; #0 $\ k |] ==> i$*k $\ j$*k" -apply (subgoal_tac "i $* intify (k) $\ j $* intify (k) ") -apply (simp (no_asm_use)) -apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) -apply (rule_tac [3] zmult_mono_lemma) -apply auto -apply (simp add: znegative_iff_zless_0 not_zless_iff_zle [THEN iff_sym]) -done - -lemma zmult_zle_mono1_neg: "[| i $\ j; k $\ #0 |] ==> j$*k $\ i$*k" -apply (rule zminus_zle_zminus [THEN iffD1]) -apply (simp del: zmult_zminus_right - add: zmult_zminus_right [symmetric] zmult_zle_mono1 zle_zminus) -done - -lemma zmult_zle_mono2: "[| i $\ j; #0 $\ k |] ==> k$*i $\ k$*j" -apply (drule zmult_zle_mono1) -apply (simp_all add: zmult_commute) -done - -lemma zmult_zle_mono2_neg: "[| i $\ j; k $\ #0 |] ==> k$*j $\ k$*i" -apply (drule zmult_zle_mono1_neg) -apply (simp_all add: zmult_commute) -done - -(* $\ monotonicity, BOTH arguments*) -lemma zmult_zle_mono: - "[| i $\ j; k $\ l; #0 $\ j; #0 $\ k |] ==> i$*k $\ j$*l" -apply (erule zmult_zle_mono1 [THEN zle_trans]) -apply assumption -apply (erule zmult_zle_mono2) -apply assumption -done - - -(** strict, in 1st argument; proof is by induction on k>0 **) - -lemma zmult_zless_mono2_lemma [rule_format]: - "[| i$ nat |] ==> 0 $#k $* i $< $#k $* j" -apply (induct_tac "k") - prefer 2 - apply (subst int_succ_int_1) - apply (erule natE) -apply (simp_all add: zadd_zmult_distrib zadd_zless_mono zle_def) -apply (frule nat_0_le) -apply (subgoal_tac "i $+ (i $+ $# xa $* i) $< j $+ (j $+ $# xa $* j) ") -apply (simp (no_asm_use)) -apply (rule zadd_zless_mono) -apply (simp_all (no_asm_simp) add: zle_def) -done - -lemma zmult_zless_mono2: "[| i$ k$*i $< k$*j" -apply (subgoal_tac "intify (k) $* i $< intify (k) $* j") -apply (simp (no_asm_use)) -apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) -apply (rule_tac [3] zmult_zless_mono2_lemma) -apply auto -apply (simp add: znegative_iff_zless_0) -apply (drule zless_trans, assumption) -apply (auto simp add: zero_lt_zmagnitude) -done - -lemma zmult_zless_mono1: "[| i$ i$*k $< j$*k" -apply (drule zmult_zless_mono2) -apply (simp_all add: zmult_commute) -done - -(* < monotonicity, BOTH arguments*) -lemma zmult_zless_mono: - "[| i $< j; k $< l; #0 $< j; #0 $< k |] ==> i$*k $< j$*l" -apply (erule zmult_zless_mono1 [THEN zless_trans]) -apply assumption -apply (erule zmult_zless_mono2) -apply assumption -done - -lemma zmult_zless_mono1_neg: "[| i $< j; k $< #0 |] ==> j$*k $< i$*k" -apply (rule zminus_zless_zminus [THEN iffD1]) -apply (simp del: zmult_zminus_right - add: zmult_zminus_right [symmetric] zmult_zless_mono1 zless_zminus) -done - -lemma zmult_zless_mono2_neg: "[| i $< j; k $< #0 |] ==> k$*j $< k$*i" -apply (rule zminus_zless_zminus [THEN iffD1]) -apply (simp del: zmult_zminus - add: zmult_zminus [symmetric] zmult_zless_mono2 zless_zminus) -done - - -(** Products of zeroes **) - -lemma zmult_eq_lemma: - "[| m \ int; n \ int |] ==> (m = #0 | n = #0) \ (m$*n = #0)" -apply (case_tac "m $< #0") -apply (auto simp add: not_zless_iff_zle zle_def neq_iff_zless) -apply (force dest: zmult_zless_mono1_neg zmult_zless_mono1)+ -done - -lemma zmult_eq_0_iff [iff]: "(m$*n = #0) \ (intify(m) = #0 | intify(n) = #0)" -apply (simp add: zmult_eq_lemma) -done - - -(** Cancellation laws for k*m < k*n and m*k < n*k, also for @{text"\"} and =, - but not (yet?) for k*m < n*k. **) - -lemma zmult_zless_lemma: - "[| k \ int; m \ int; n \ int |] - ==> (m$*k $< n$*k) \ ((#0 $< k & m$ ((#0 $< k & m$ ((#0 $< k & m$ n$*k) \ ((#0 $< k \ m$\n) & (k $< #0 \ n$\m))" -by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel2) - -lemma zmult_zle_cancel1: - "(k$*m $\ k$*n) \ ((#0 $< k \ m$\n) & (k $< #0 \ n$\m))" -by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel1) - -lemma int_eq_iff_zle: "[| m \ int; n \ int |] ==> m=n \ (m $\ n & n $\ m)" -apply (blast intro: zle_refl zle_anti_sym) -done - -lemma zmult_cancel2_lemma: - "[| k \ int; m \ int; n \ int |] ==> (m$*k = n$*k) \ (k=#0 | m=n)" -apply (simp add: int_eq_iff_zle [of "m$*k"] int_eq_iff_zle [of m]) -apply (auto simp add: zmult_zle_cancel2 neq_iff_zless) -done - -lemma zmult_cancel2 [simp]: - "(m$*k = n$*k) \ (intify(k) = #0 | intify(m) = intify(n))" -apply (rule iff_trans) -apply (rule_tac [2] zmult_cancel2_lemma) -apply auto -done - -lemma zmult_cancel1 [simp]: - "(k$*m = k$*n) \ (intify(k) = #0 | intify(m) = intify(n))" -by (simp add: zmult_commute [of k] zmult_cancel2) - - -subsection\Uniqueness and monotonicity of quotients and remainders\ - -lemma unique_quotient_lemma: - "[| b$*q' $+ r' $\ b$*q $+ r; #0 $\ r'; #0 $< b; r $< b |] - ==> q' $\ q" -apply (subgoal_tac "r' $+ b $* (q'$-q) $\ r") - prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_ac zcompare_rls) -apply (subgoal_tac "#0 $< b $* (#1 $+ q $- q') ") - prefer 2 - apply (erule zle_zless_trans) - apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls) - apply (erule zle_zless_trans) - apply simp -apply (subgoal_tac "b $* q' $< b $* (#1 $+ q)") - prefer 2 - apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls) -apply (auto elim: zless_asym - simp add: zmult_zless_cancel1 zless_add1_iff_zle zadd_ac zcompare_rls) -done - -lemma unique_quotient_lemma_neg: - "[| b$*q' $+ r' $\ b$*q $+ r; r $\ #0; b $< #0; b $< r' |] - ==> q $\ q'" -apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r" - in unique_quotient_lemma) -apply (auto simp del: zminus_zadd_distrib - simp add: zminus_zadd_distrib [symmetric] zle_zminus zless_zminus) -done - - -lemma unique_quotient: - "[| quorem (, ); quorem (, ); b \ int; b \ #0; - q \ int; q' \ int |] ==> q = q'" -apply (simp add: split_ifs quorem_def neq_iff_zless) -apply safe -apply simp_all -apply (blast intro: zle_anti_sym - dest: zle_eq_refl [THEN unique_quotient_lemma] - zle_eq_refl [THEN unique_quotient_lemma_neg] sym)+ -done - -lemma unique_remainder: - "[| quorem (, ); quorem (, ); b \ int; b \ #0; - q \ int; q' \ int; - r \ int; r' \ int |] ==> r = r'" -apply (subgoal_tac "q = q'") - prefer 2 apply (blast intro: unique_quotient) -apply (simp add: quorem_def) -done - - -subsection\Correctness of posDivAlg, - the Division Algorithm for \a\0\ and \b>0\\ - -lemma adjust_eq [simp]: - "adjust(b, ) = (let diff = r$-b in - if #0 $\ diff then <#2$*q $+ #1,diff> - else <#2$*q,r>)" -by (simp add: Let_def adjust_def) - - -lemma posDivAlg_termination: - "[| #0 $< b; ~ a $< b |] - ==> nat_of(a $- #2 $* b $+ #1) < nat_of(a $- b $+ #1)" -apply (simp (no_asm) add: zless_nat_conj) -apply (simp add: not_zless_iff_zle zless_add1_iff_zle zcompare_rls) -done - -lemmas posDivAlg_unfold = def_wfrec [OF posDivAlg_def wf_measure] - -lemma posDivAlg_eqn: - "[| #0 $< b; a \ int; b \ int |] ==> - posDivAlg() = - (if a$ else adjust(b, posDivAlg ()))" -apply (rule posDivAlg_unfold [THEN trans]) -apply (simp add: vimage_iff not_zless_iff_zle [THEN iff_sym]) -apply (blast intro: posDivAlg_termination) -done - -lemma posDivAlg_induct_lemma [rule_format]: - assumes prem: - "!!a b. [| a \ int; b \ int; - ~ (a $< b | b $\ #0) \ P() |] ==> P()" - shows " \ int*int \ P()" -using wf_measure [where A = "int*int" and f = "%.nat_of (a $- b $+ #1)"] -proof (induct "" arbitrary: u v rule: wf_induct) - case (step x) - hence uv: "u \ int" "v \ int" by auto - thus ?case - apply (rule prem) - apply (rule impI) - apply (rule step) - apply (auto simp add: step uv not_zle_iff_zless posDivAlg_termination) - done -qed - - -lemma posDivAlg_induct [consumes 2]: - assumes u_int: "u \ int" - and v_int: "v \ int" - and ih: "!!a b. [| a \ int; b \ int; - ~ (a $< b | b $\ #0) \ P(a, #2 $* b) |] ==> P(a,b)" - shows "P(u,v)" -apply (subgoal_tac "(%. P (x,y)) ()") -apply simp -apply (rule posDivAlg_induct_lemma) -apply (simp (no_asm_use)) -apply (rule ih) -apply (auto simp add: u_int v_int) -done - -(*FIXME: use intify in integ_of so that we always have @{term"integ_of w \ int"}. - then this rewrite can work for all constants!!*) -lemma intify_eq_0_iff_zle: "intify(m) = #0 \ (m $\ #0 & #0 $\ m)" - by (simp add: int_eq_iff_zle) - - -subsection\Some convenient biconditionals for products of signs\ - -lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j" - by (drule zmult_zless_mono1, auto) - -lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j" - by (drule zmult_zless_mono1_neg, auto) - -lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0" - by (drule zmult_zless_mono1_neg, auto) - - -(** Inequality reasoning **) - -lemma int_0_less_lemma: - "[| x \ int; y \ int |] - ==> (#0 $< x $* y) \ (#0 $< x & #0 $< y | x $< #0 & y $< #0)" -apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg) -apply (rule ccontr) -apply (rule_tac [2] ccontr) -apply (auto simp add: zle_def not_zless_iff_zle) -apply (erule_tac P = "#0$< x$* y" in rev_mp) -apply (erule_tac [2] P = "#0$< x$* y" in rev_mp) -apply (drule zmult_pos_neg, assumption) - prefer 2 - apply (drule zmult_pos_neg, assumption) -apply (auto dest: zless_not_sym simp add: zmult_commute) -done - -lemma int_0_less_mult_iff: - "(#0 $< x $* y) \ (#0 $< x & #0 $< y | x $< #0 & y $< #0)" -apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_less_lemma) -apply auto -done - -lemma int_0_le_lemma: - "[| x \ int; y \ int |] - ==> (#0 $\ x $* y) \ (#0 $\ x & #0 $\ y | x $\ #0 & y $\ #0)" -by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff) - -lemma int_0_le_mult_iff: - "(#0 $\ x $* y) \ ((#0 $\ x & #0 $\ y) | (x $\ #0 & y $\ #0))" -apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_le_lemma) -apply auto -done - -lemma zmult_less_0_iff: - "(x $* y $< #0) \ (#0 $< x & y $< #0 | x $< #0 & #0 $< y)" -apply (auto simp add: int_0_le_mult_iff not_zle_iff_zless [THEN iff_sym]) -apply (auto dest: zless_not_sym simp add: not_zle_iff_zless) -done - -lemma zmult_le_0_iff: - "(x $* y $\ #0) \ (#0 $\ x & y $\ #0 | x $\ #0 & #0 $\ y)" -by (auto dest: zless_not_sym - simp add: int_0_less_mult_iff not_zless_iff_zle [THEN iff_sym]) - - -(*Typechecking for posDivAlg*) -lemma posDivAlg_type [rule_format]: - "[| a \ int; b \ int |] ==> posDivAlg() \ int * int" -apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) -apply assumption+ -apply (case_tac "#0 $< ba") - apply (simp add: posDivAlg_eqn adjust_def integ_of_type - split: split_if_asm) - apply clarify - apply (simp add: int_0_less_mult_iff not_zle_iff_zless) -apply (simp add: not_zless_iff_zle) -apply (subst posDivAlg_unfold) -apply simp -done - -(*Correctness of posDivAlg: it computes quotients correctly*) -lemma posDivAlg_correct [rule_format]: - "[| a \ int; b \ int |] - ==> #0 $\ a \ #0 $< b \ quorem (, posDivAlg())" -apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) -apply auto - apply (simp_all add: quorem_def) - txt\base case: a - apply (simp add: posDivAlg_eqn) - apply (simp add: not_zless_iff_zle [THEN iff_sym]) - apply (simp add: int_0_less_mult_iff) -txt\main argument\ -apply (subst posDivAlg_eqn) -apply (simp_all (no_asm_simp)) -apply (erule splitE) -apply (rule posDivAlg_type) -apply (simp_all add: int_0_less_mult_iff) -apply (auto simp add: zadd_zmult_distrib2 Let_def) -txt\now just linear arithmetic\ -apply (simp add: not_zle_iff_zless zdiff_zless_iff) -done - - -subsection\Correctness of negDivAlg, the division algorithm for a<0 and b>0\ - -lemma negDivAlg_termination: - "[| #0 $< b; a $+ b $< #0 |] - ==> nat_of($- a $- #2 $* b) < nat_of($- a $- b)" -apply (simp (no_asm) add: zless_nat_conj) -apply (simp add: zcompare_rls not_zle_iff_zless zless_zdiff_iff [THEN iff_sym] - zless_zminus) -done - -lemmas negDivAlg_unfold = def_wfrec [OF negDivAlg_def wf_measure] - -lemma negDivAlg_eqn: - "[| #0 $< b; a \ int; b \ int |] ==> - negDivAlg() = - (if #0 $\ a$+b then <#-1,a$+b> - else adjust(b, negDivAlg ()))" -apply (rule negDivAlg_unfold [THEN trans]) -apply (simp (no_asm_simp) add: vimage_iff not_zless_iff_zle [THEN iff_sym]) -apply (blast intro: negDivAlg_termination) -done - -lemma negDivAlg_induct_lemma [rule_format]: - assumes prem: - "!!a b. [| a \ int; b \ int; - ~ (#0 $\ a $+ b | b $\ #0) \ P() |] - ==> P()" - shows " \ int*int \ P()" -using wf_measure [where A = "int*int" and f = "%.nat_of ($- a $- b)"] -proof (induct "" arbitrary: u v rule: wf_induct) - case (step x) - hence uv: "u \ int" "v \ int" by auto - thus ?case - apply (rule prem) - apply (rule impI) - apply (rule step) - apply (auto simp add: step uv not_zle_iff_zless negDivAlg_termination) - done -qed - -lemma negDivAlg_induct [consumes 2]: - assumes u_int: "u \ int" - and v_int: "v \ int" - and ih: "!!a b. [| a \ int; b \ int; - ~ (#0 $\ a $+ b | b $\ #0) \ P(a, #2 $* b) |] - ==> P(a,b)" - shows "P(u,v)" -apply (subgoal_tac " (%. P (x,y)) ()") -apply simp -apply (rule negDivAlg_induct_lemma) -apply (simp (no_asm_use)) -apply (rule ih) -apply (auto simp add: u_int v_int) -done - - -(*Typechecking for negDivAlg*) -lemma negDivAlg_type: - "[| a \ int; b \ int |] ==> negDivAlg() \ int * int" -apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) -apply assumption+ -apply (case_tac "#0 $< ba") - apply (simp add: negDivAlg_eqn adjust_def integ_of_type - split: split_if_asm) - apply clarify - apply (simp add: int_0_less_mult_iff not_zle_iff_zless) -apply (simp add: not_zless_iff_zle) -apply (subst negDivAlg_unfold) -apply simp -done - - -(*Correctness of negDivAlg: it computes quotients correctly - It doesn't work if a=0 because the 0/b=0 rather than -1*) -lemma negDivAlg_correct [rule_format]: - "[| a \ int; b \ int |] - ==> a $< #0 \ #0 $< b \ quorem (, negDivAlg())" -apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) - apply auto - apply (simp_all add: quorem_def) - txt\base case: @{term "0$\a$+b"}\ - apply (simp add: negDivAlg_eqn) - apply (simp add: not_zless_iff_zle [THEN iff_sym]) - apply (simp add: int_0_less_mult_iff) -txt\main argument\ -apply (subst negDivAlg_eqn) -apply (simp_all (no_asm_simp)) -apply (erule splitE) -apply (rule negDivAlg_type) -apply (simp_all add: int_0_less_mult_iff) -apply (auto simp add: zadd_zmult_distrib2 Let_def) -txt\now just linear arithmetic\ -apply (simp add: not_zle_iff_zless zdiff_zless_iff) -done - - -subsection\Existence shown by proving the division algorithm to be correct\ - -(*the case a=0*) -lemma quorem_0: "[|b \ #0; b \ int|] ==> quorem (<#0,b>, <#0,#0>)" -by (force simp add: quorem_def neq_iff_zless) - -lemma posDivAlg_zero_divisor: "posDivAlg() = <#0,a>" -apply (subst posDivAlg_unfold) -apply simp -done - -lemma posDivAlg_0 [simp]: "posDivAlg (<#0,b>) = <#0,#0>" -apply (subst posDivAlg_unfold) -apply (simp add: not_zle_iff_zless) -done - - -(*Needed below. Actually it's an equivalence.*) -lemma linear_arith_lemma: "~ (#0 $\ #-1 $+ b) ==> (b $\ #0)" -apply (simp add: not_zle_iff_zless) -apply (drule zminus_zless_zminus [THEN iffD2]) -apply (simp add: zadd_commute zless_add1_iff_zle zle_zminus) -done - -lemma negDivAlg_minus1 [simp]: "negDivAlg (<#-1,b>) = <#-1, b$-#1>" -apply (subst negDivAlg_unfold) -apply (simp add: linear_arith_lemma integ_of_type vimage_iff) -done - -lemma negateSnd_eq [simp]: "negateSnd () = " -apply (unfold negateSnd_def) -apply auto -done - -lemma negateSnd_type: "qr \ int * int ==> negateSnd (qr) \ int * int" -apply (unfold negateSnd_def) -apply auto -done - -lemma quorem_neg: - "[|quorem (<$-a,$-b>, qr); a \ int; b \ int; qr \ int * int|] - ==> quorem (, negateSnd(qr))" -apply clarify -apply (auto elim: zless_asym simp add: quorem_def zless_zminus) -txt\linear arithmetic from here on\ -apply (simp_all add: zminus_equation [of a] zminus_zless) -apply (cut_tac [2] z = "b" and w = "#0" in zless_linear) -apply (cut_tac [1] z = "b" and w = "#0" in zless_linear) -apply auto -apply (blast dest: zle_zless_trans)+ -done - -lemma divAlg_correct: - "[|b \ #0; a \ int; b \ int|] ==> quorem (, divAlg())" -apply (auto simp add: quorem_0 divAlg_def) -apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct - posDivAlg_type negDivAlg_type) -apply (auto simp add: quorem_def neq_iff_zless) -txt\linear arithmetic from here on\ -apply (auto simp add: zle_def) -done - -lemma divAlg_type: "[|a \ int; b \ int|] ==> divAlg() \ int * int" -apply (auto simp add: divAlg_def) -apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type) -done - - -(** intify cancellation **) - -lemma zdiv_intify1 [simp]: "intify(x) zdiv y = x zdiv y" - by (simp add: zdiv_def) - -lemma zdiv_intify2 [simp]: "x zdiv intify(y) = x zdiv y" - by (simp add: zdiv_def) - -lemma zdiv_type [iff,TC]: "z zdiv w \ int" -apply (unfold zdiv_def) -apply (blast intro: fst_type divAlg_type) -done - -lemma zmod_intify1 [simp]: "intify(x) zmod y = x zmod y" - by (simp add: zmod_def) - -lemma zmod_intify2 [simp]: "x zmod intify(y) = x zmod y" - by (simp add: zmod_def) - -lemma zmod_type [iff,TC]: "z zmod w \ int" -apply (unfold zmod_def) -apply (rule snd_type) -apply (blast intro: divAlg_type) -done - - -(** Arbitrary definitions for division by zero. Useful to simplify - certain equations **) - -lemma DIVISION_BY_ZERO_ZDIV: "a zdiv #0 = #0" - by (simp add: zdiv_def divAlg_def posDivAlg_zero_divisor) - -lemma DIVISION_BY_ZERO_ZMOD: "a zmod #0 = intify(a)" - by (simp add: zmod_def divAlg_def posDivAlg_zero_divisor) - - -(** Basic laws about division and remainder **) - -lemma raw_zmod_zdiv_equality: - "[| a \ int; b \ int |] ==> a = b $* (a zdiv b) $+ (a zmod b)" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (cut_tac a = "a" and b = "b" in divAlg_correct) -apply (auto simp add: quorem_def zdiv_def zmod_def split_def) -done - -lemma zmod_zdiv_equality: "intify(a) = b $* (a zdiv b) $+ (a zmod b)" -apply (rule trans) -apply (rule_tac b = "intify (b)" in raw_zmod_zdiv_equality) -apply auto -done - -lemma pos_mod: "#0 $< b ==> #0 $\ a zmod b & a zmod b $< b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) -apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) -apply (blast dest: zle_zless_trans)+ -done - -lemmas pos_mod_sign = pos_mod [THEN conjunct1] - and pos_mod_bound = pos_mod [THEN conjunct2] - -lemma neg_mod: "b $< #0 ==> a zmod b $\ #0 & b $< a zmod b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) -apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) -apply (blast dest: zle_zless_trans) -apply (blast dest: zless_trans)+ -done - -lemmas neg_mod_sign = neg_mod [THEN conjunct1] - and neg_mod_bound = neg_mod [THEN conjunct2] - - -(** proving general properties of zdiv and zmod **) - -lemma quorem_div_mod: - "[|b \ #0; a \ int; b \ int |] - ==> quorem (, )" -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) -apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound - neg_mod_sign neg_mod_bound) -done - -(*Surely quorem(,) implies @{term"a \ int"}, but it doesn't matter*) -lemma quorem_div: - "[| quorem(,); b \ #0; a \ int; b \ int; q \ int |] - ==> a zdiv b = q" -by (blast intro: quorem_div_mod [THEN unique_quotient]) - -lemma quorem_mod: - "[| quorem(,); b \ #0; a \ int; b \ int; q \ int; r \ int |] - ==> a zmod b = r" -by (blast intro: quorem_div_mod [THEN unique_remainder]) - -lemma zdiv_pos_pos_trivial_raw: - "[| a \ int; b \ int; #0 $\ a; a $< b |] ==> a zdiv b = #0" -apply (rule quorem_div) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zle_zless_trans)+ -done - -lemma zdiv_pos_pos_trivial: "[| #0 $\ a; a $< b |] ==> a zdiv b = #0" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zdiv_pos_pos_trivial_raw) -apply auto -done - -lemma zdiv_neg_neg_trivial_raw: - "[| a \ int; b \ int; a $\ #0; b $< a |] ==> a zdiv b = #0" -apply (rule_tac r = "a" in quorem_div) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zle_zless_trans zless_trans)+ -done - -lemma zdiv_neg_neg_trivial: "[| a $\ #0; b $< a |] ==> a zdiv b = #0" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zdiv_neg_neg_trivial_raw) -apply auto -done - -lemma zadd_le_0_lemma: "[| a$+b $\ #0; #0 $< a; #0 $< b |] ==> False" -apply (drule_tac z' = "#0" and z = "b" in zadd_zless_mono) -apply (auto simp add: zle_def) -apply (blast dest: zless_trans) -done - -lemma zdiv_pos_neg_trivial_raw: - "[| a \ int; b \ int; #0 $< a; a$+b $\ #0 |] ==> a zdiv b = #-1" -apply (rule_tac r = "a $+ b" in quorem_div) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ -done - -lemma zdiv_pos_neg_trivial: "[| #0 $< a; a$+b $\ #0 |] ==> a zdiv b = #-1" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zdiv_pos_neg_trivial_raw) -apply auto -done - -(*There is no zdiv_neg_pos_trivial because #0 zdiv b = #0 would supersede it*) - - -lemma zmod_pos_pos_trivial_raw: - "[| a \ int; b \ int; #0 $\ a; a $< b |] ==> a zmod b = a" -apply (rule_tac q = "#0" in quorem_mod) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zle_zless_trans)+ -done - -lemma zmod_pos_pos_trivial: "[| #0 $\ a; a $< b |] ==> a zmod b = intify(a)" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zmod_pos_pos_trivial_raw) -apply auto -done - -lemma zmod_neg_neg_trivial_raw: - "[| a \ int; b \ int; a $\ #0; b $< a |] ==> a zmod b = a" -apply (rule_tac q = "#0" in quorem_mod) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zle_zless_trans zless_trans)+ -done - -lemma zmod_neg_neg_trivial: "[| a $\ #0; b $< a |] ==> a zmod b = intify(a)" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zmod_neg_neg_trivial_raw) -apply auto -done - -lemma zmod_pos_neg_trivial_raw: - "[| a \ int; b \ int; #0 $< a; a$+b $\ #0 |] ==> a zmod b = a$+b" -apply (rule_tac q = "#-1" in quorem_mod) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ -done - -lemma zmod_pos_neg_trivial: "[| #0 $< a; a$+b $\ #0 |] ==> a zmod b = a$+b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zmod_pos_neg_trivial_raw) -apply auto -done - -(*There is no zmod_neg_pos_trivial...*) - - -(*Simpler laws such as -a zdiv b = -(a zdiv b) FAIL*) - -lemma zdiv_zminus_zminus_raw: - "[|a \ int; b \ int|] ==> ($-a) zdiv ($-b) = a zdiv b" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_div]) -apply auto -done - -lemma zdiv_zminus_zminus [simp]: "($-a) zdiv ($-b) = a zdiv b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zminus_zminus_raw) -apply auto -done - -(*Simpler laws such as -a zmod b = -(a zmod b) FAIL*) -lemma zmod_zminus_zminus_raw: - "[|a \ int; b \ int|] ==> ($-a) zmod ($-b) = $- (a zmod b)" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod]) -apply auto -done - -lemma zmod_zminus_zminus [simp]: "($-a) zmod ($-b) = $- (a zmod b)" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zminus_zminus_raw) -apply auto -done - - -subsection\division of a number by itself\ - -lemma self_quotient_aux1: "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $\ q" -apply (subgoal_tac "#0 $< a$*q") -apply (cut_tac w = "#0" and z = "q" in add1_zle_iff) -apply (simp add: int_0_less_mult_iff) -apply (blast dest: zless_trans) -(*linear arithmetic...*) -apply (drule_tac t = "%x. x $- r" in subst_context) -apply (drule sym) -apply (simp add: zcompare_rls) -done - -lemma self_quotient_aux2: "[| #0 $< a; a = r $+ a$*q; #0 $\ r |] ==> q $\ #1" -apply (subgoal_tac "#0 $\ a$* (#1$-q)") - apply (simp add: int_0_le_mult_iff zcompare_rls) - apply (blast dest: zle_zless_trans) -apply (simp add: zdiff_zmult_distrib2) -apply (drule_tac t = "%x. x $- a $* q" in subst_context) -apply (simp add: zcompare_rls) -done - -lemma self_quotient: - "[| quorem(,); a \ int; q \ int; a \ #0|] ==> q = #1" -apply (simp add: split_ifs quorem_def neq_iff_zless) -apply (rule zle_anti_sym) -apply safe -apply auto -prefer 4 apply (blast dest: zless_trans) -apply (blast dest: zless_trans) -apply (rule_tac [3] a = "$-a" and r = "$-r" in self_quotient_aux1) -apply (rule_tac a = "$-a" and r = "$-r" in self_quotient_aux2) -apply (rule_tac [6] zminus_equation [THEN iffD1]) -apply (rule_tac [2] zminus_equation [THEN iffD1]) -apply (force intro: self_quotient_aux1 self_quotient_aux2 - simp add: zadd_commute zmult_zminus)+ -done - -lemma self_remainder: - "[|quorem(,); a \ int; q \ int; r \ int; a \ #0|] ==> r = #0" -apply (frule self_quotient) -apply (auto simp add: quorem_def) -done - -lemma zdiv_self_raw: "[|a \ #0; a \ int|] ==> a zdiv a = #1" -apply (blast intro: quorem_div_mod [THEN self_quotient]) -done - -lemma zdiv_self [simp]: "intify(a) \ #0 ==> a zdiv a = #1" -apply (drule zdiv_self_raw) -apply auto -done - -(*Here we have 0 zmod 0 = 0, also assumed by Knuth (who puts m zmod 0 = 0) *) -lemma zmod_self_raw: "a \ int ==> a zmod a = #0" -apply (case_tac "a = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (blast intro: quorem_div_mod [THEN self_remainder]) -done - -lemma zmod_self [simp]: "a zmod a = #0" -apply (cut_tac a = "intify (a)" in zmod_self_raw) -apply auto -done - - -subsection\Computation of division and remainder\ - -lemma zdiv_zero [simp]: "#0 zdiv b = #0" - by (simp add: zdiv_def divAlg_def) - -lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1" - by (simp (no_asm_simp) add: zdiv_def divAlg_def) - -lemma zmod_zero [simp]: "#0 zmod b = #0" - by (simp add: zmod_def divAlg_def) - -lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1" - by (simp add: zdiv_def divAlg_def) - -lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1" - by (simp add: zmod_def divAlg_def) - -(** a positive, b positive **) - -lemma zdiv_pos_pos: "[| #0 $< a; #0 $\ b |] - ==> a zdiv b = fst (posDivAlg())" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def) -apply (auto simp add: zle_def) -done - -lemma zmod_pos_pos: - "[| #0 $< a; #0 $\ b |] - ==> a zmod b = snd (posDivAlg())" -apply (simp (no_asm_simp) add: zmod_def divAlg_def) -apply (auto simp add: zle_def) -done - -(** a negative, b positive **) - -lemma zdiv_neg_pos: - "[| a $< #0; #0 $< b |] - ==> a zdiv b = fst (negDivAlg())" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def) -apply (blast dest: zle_zless_trans) -done - -lemma zmod_neg_pos: - "[| a $< #0; #0 $< b |] - ==> a zmod b = snd (negDivAlg())" -apply (simp (no_asm_simp) add: zmod_def divAlg_def) -apply (blast dest: zle_zless_trans) -done - -(** a positive, b negative **) - -lemma zdiv_pos_neg: - "[| #0 $< a; b $< #0 |] - ==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def intify_eq_0_iff_zle) -apply auto -apply (blast dest: zle_zless_trans)+ -apply (blast dest: zless_trans) -apply (blast intro: zless_imp_zle) -done - -lemma zmod_pos_neg: - "[| #0 $< a; b $< #0 |] - ==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))" -apply (simp (no_asm_simp) add: zmod_def divAlg_def intify_eq_0_iff_zle) -apply auto -apply (blast dest: zle_zless_trans)+ -apply (blast dest: zless_trans) -apply (blast intro: zless_imp_zle) -done - -(** a negative, b negative **) - -lemma zdiv_neg_neg: - "[| a $< #0; b $\ #0 |] - ==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def) -apply auto -apply (blast dest!: zle_zless_trans)+ -done - -lemma zmod_neg_neg: - "[| a $< #0; b $\ #0 |] - ==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))" -apply (simp (no_asm_simp) add: zmod_def divAlg_def) -apply auto -apply (blast dest!: zle_zless_trans)+ -done - -declare zdiv_pos_pos [of "integ_of (v)" "integ_of (w)", simp] for v w -declare zdiv_neg_pos [of "integ_of (v)" "integ_of (w)", simp] for v w -declare zdiv_pos_neg [of "integ_of (v)" "integ_of (w)", simp] for v w -declare zdiv_neg_neg [of "integ_of (v)" "integ_of (w)", simp] for v w -declare zmod_pos_pos [of "integ_of (v)" "integ_of (w)", simp] for v w -declare zmod_neg_pos [of "integ_of (v)" "integ_of (w)", simp] for v w -declare zmod_pos_neg [of "integ_of (v)" "integ_of (w)", simp] for v w -declare zmod_neg_neg [of "integ_of (v)" "integ_of (w)", simp] for v w -declare posDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", simp] for v w -declare negDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", simp] for v w - - -(** Special-case simplification **) - -lemma zmod_1 [simp]: "a zmod #1 = #0" -apply (cut_tac a = "a" and b = "#1" in pos_mod_sign) -apply (cut_tac [2] a = "a" and b = "#1" in pos_mod_bound) -apply auto -(*arithmetic*) -apply (drule add1_zle_iff [THEN iffD2]) -apply (rule zle_anti_sym) -apply auto -done - -lemma zdiv_1 [simp]: "a zdiv #1 = intify(a)" -apply (cut_tac a = "a" and b = "#1" in zmod_zdiv_equality) -apply auto -done - -lemma zmod_minus1_right [simp]: "a zmod #-1 = #0" -apply (cut_tac a = "a" and b = "#-1" in neg_mod_sign) -apply (cut_tac [2] a = "a" and b = "#-1" in neg_mod_bound) -apply auto -(*arithmetic*) -apply (drule add1_zle_iff [THEN iffD2]) -apply (rule zle_anti_sym) -apply auto -done - -lemma zdiv_minus1_right_raw: "a \ int ==> a zdiv #-1 = $-a" -apply (cut_tac a = "a" and b = "#-1" in zmod_zdiv_equality) -apply auto -apply (rule equation_zminus [THEN iffD2]) -apply auto -done - -lemma zdiv_minus1_right: "a zdiv #-1 = $-a" -apply (cut_tac a = "intify (a)" in zdiv_minus1_right_raw) -apply auto -done -declare zdiv_minus1_right [simp] - - -subsection\Monotonicity in the first argument (divisor)\ - -lemma zdiv_mono1: "[| a $\ a'; #0 $< b |] ==> a zdiv b $\ a' zdiv b" -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) -apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) -apply (rule unique_quotient_lemma) -apply (erule subst) -apply (erule subst) -apply (simp_all (no_asm_simp) add: pos_mod_sign pos_mod_bound) -done - -lemma zdiv_mono1_neg: "[| a $\ a'; b $< #0 |] ==> a' zdiv b $\ a zdiv b" -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) -apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) -apply (rule unique_quotient_lemma_neg) -apply (erule subst) -apply (erule subst) -apply (simp_all (no_asm_simp) add: neg_mod_sign neg_mod_bound) -done - - -subsection\Monotonicity in the second argument (dividend)\ - -lemma q_pos_lemma: - "[| #0 $\ b'$*q' $+ r'; r' $< b'; #0 $< b' |] ==> #0 $\ q'" -apply (subgoal_tac "#0 $< b'$* (q' $+ #1)") - apply (simp add: int_0_less_mult_iff) - apply (blast dest: zless_trans intro: zless_add1_iff_zle [THEN iffD1]) -apply (simp add: zadd_zmult_distrib2) -apply (erule zle_zless_trans) -apply (erule zadd_zless_mono2) -done - -lemma zdiv_mono2_lemma: - "[| b$*q $+ r = b'$*q' $+ r'; #0 $\ b'$*q' $+ r'; - r' $< b'; #0 $\ r; #0 $< b'; b' $\ b |] - ==> q $\ q'" -apply (frule q_pos_lemma, assumption+) -apply (subgoal_tac "b$*q $< b$* (q' $+ #1)") - apply (simp add: zmult_zless_cancel1) - apply (force dest: zless_add1_iff_zle [THEN iffD1] zless_trans zless_zle_trans) -apply (subgoal_tac "b$*q = r' $- r $+ b'$*q'") - prefer 2 apply (simp add: zcompare_rls) -apply (simp (no_asm_simp) add: zadd_zmult_distrib2) -apply (subst zadd_commute [of "b $* q'"], rule zadd_zless_mono) - prefer 2 apply (blast intro: zmult_zle_mono1) -apply (subgoal_tac "r' $+ #0 $< b $+ r") - apply (simp add: zcompare_rls) -apply (rule zadd_zless_mono) - apply auto -apply (blast dest: zless_zle_trans) -done - - -lemma zdiv_mono2_raw: - "[| #0 $\ a; #0 $< b'; b' $\ b; a \ int |] - ==> a zdiv b $\ a zdiv b'" -apply (subgoal_tac "#0 $< b") - prefer 2 apply (blast dest: zless_zle_trans) -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) -apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality) -apply (rule zdiv_mono2_lemma) -apply (erule subst) -apply (erule subst) -apply (simp_all add: pos_mod_sign pos_mod_bound) -done - -lemma zdiv_mono2: - "[| #0 $\ a; #0 $< b'; b' $\ b |] - ==> a zdiv b $\ a zdiv b'" -apply (cut_tac a = "intify (a)" in zdiv_mono2_raw) -apply auto -done - -lemma q_neg_lemma: - "[| b'$*q' $+ r' $< #0; #0 $\ r'; #0 $< b' |] ==> q' $< #0" -apply (subgoal_tac "b'$*q' $< #0") - prefer 2 apply (force intro: zle_zless_trans) -apply (simp add: zmult_less_0_iff) -apply (blast dest: zless_trans) -done - - - -lemma zdiv_mono2_neg_lemma: - "[| b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0; - r $< b; #0 $\ r'; #0 $< b'; b' $\ b |] - ==> q' $\ q" -apply (subgoal_tac "#0 $< b") - prefer 2 apply (blast dest: zless_zle_trans) -apply (frule q_neg_lemma, assumption+) -apply (subgoal_tac "b$*q' $< b$* (q $+ #1)") - apply (simp add: zmult_zless_cancel1) - apply (blast dest: zless_trans zless_add1_iff_zle [THEN iffD1]) -apply (simp (no_asm_simp) add: zadd_zmult_distrib2) -apply (subgoal_tac "b$*q' $\ b'$*q'") - prefer 2 - apply (simp add: zmult_zle_cancel2) - apply (blast dest: zless_trans) -apply (subgoal_tac "b'$*q' $+ r $< b $+ (b$*q $+ r)") - prefer 2 - apply (erule ssubst) - apply simp - apply (drule_tac w' = "r" and z' = "#0" in zadd_zless_mono) - apply (assumption) - apply simp -apply (simp (no_asm_use) add: zadd_commute) -apply (rule zle_zless_trans) - prefer 2 apply (assumption) -apply (simp (no_asm_simp) add: zmult_zle_cancel2) -apply (blast dest: zless_trans) -done - -lemma zdiv_mono2_neg_raw: - "[| a $< #0; #0 $< b'; b' $\ b; a \ int |] - ==> a zdiv b' $\ a zdiv b" -apply (subgoal_tac "#0 $< b") - prefer 2 apply (blast dest: zless_zle_trans) -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) -apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality) -apply (rule zdiv_mono2_neg_lemma) -apply (erule subst) -apply (erule subst) -apply (simp_all add: pos_mod_sign pos_mod_bound) -done - -lemma zdiv_mono2_neg: "[| a $< #0; #0 $< b'; b' $\ b |] - ==> a zdiv b' $\ a zdiv b" -apply (cut_tac a = "intify (a)" in zdiv_mono2_neg_raw) -apply auto -done - - - -subsection\More algebraic laws for zdiv and zmod\ - -(** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **) - -lemma zmult1_lemma: - "[| quorem(, ); c \ int; c \ #0 |] - ==> quorem (, )" -apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 - pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) -apply (auto intro: raw_zmod_zdiv_equality) -done - -lemma zdiv_zmult1_eq_raw: - "[|b \ int; c \ int|] - ==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" -apply (case_tac "c = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_div]) -apply auto -done - -lemma zdiv_zmult1_eq: "(a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" -apply (cut_tac b = "intify (b)" and c = "intify (c)" in zdiv_zmult1_eq_raw) -apply auto -done - -lemma zmod_zmult1_eq_raw: - "[|b \ int; c \ int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c" -apply (case_tac "c = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod]) -apply auto -done - -lemma zmod_zmult1_eq: "(a$*b) zmod c = a$*(b zmod c) zmod c" -apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult1_eq_raw) -apply auto -done - -lemma zmod_zmult1_eq': "(a$*b) zmod c = ((a zmod c) $* b) zmod c" -apply (rule trans) -apply (rule_tac b = " (b $* a) zmod c" in trans) -apply (rule_tac [2] zmod_zmult1_eq) -apply (simp_all (no_asm) add: zmult_commute) -done - -lemma zmod_zmult_distrib: "(a$*b) zmod c = ((a zmod c) $* (b zmod c)) zmod c" -apply (rule zmod_zmult1_eq' [THEN trans]) -apply (rule zmod_zmult1_eq) -done - -lemma zdiv_zmult_self1 [simp]: "intify(b) \ #0 ==> (a$*b) zdiv b = intify(a)" - by (simp add: zdiv_zmult1_eq) - -lemma zdiv_zmult_self2 [simp]: "intify(b) \ #0 ==> (b$*a) zdiv b = intify(a)" - by (simp add: zmult_commute) - -lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0" - by (simp add: zmod_zmult1_eq) - -lemma zmod_zmult_self2 [simp]: "(b$*a) zmod b = #0" - by (simp add: zmult_commute zmod_zmult1_eq) - - -(** proving (a$+b) zdiv c = - a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **) - -lemma zadd1_lemma: - "[| quorem(, ); quorem(, ); - c \ int; c \ #0 |] - ==> quorem (, )" -apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 - pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) -apply (auto intro: raw_zmod_zdiv_equality) -done - -(*NOT suitable for rewriting: the RHS has an instance of the LHS*) -lemma zdiv_zadd1_eq_raw: - "[|a \ int; b \ int; c \ int|] ==> - (a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" -apply (case_tac "c = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, - THEN quorem_div]) -done - -lemma zdiv_zadd1_eq: - "(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" -apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" - in zdiv_zadd1_eq_raw) -apply auto -done - -lemma zmod_zadd1_eq_raw: - "[|a \ int; b \ int; c \ int|] - ==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" -apply (case_tac "c = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, - THEN quorem_mod]) -done - -lemma zmod_zadd1_eq: "(a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" -apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" - in zmod_zadd1_eq_raw) -apply auto -done - -lemma zmod_div_trivial_raw: - "[|a \ int; b \ int|] ==> (a zmod b) zdiv b = #0" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound - zdiv_pos_pos_trivial neg_mod_sign neg_mod_bound zdiv_neg_neg_trivial) -done - -lemma zmod_div_trivial [simp]: "(a zmod b) zdiv b = #0" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_div_trivial_raw) -apply auto -done - -lemma zmod_mod_trivial_raw: - "[|a \ int; b \ int|] ==> (a zmod b) zmod b = a zmod b" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound - zmod_pos_pos_trivial neg_mod_sign neg_mod_bound zmod_neg_neg_trivial) -done - -lemma zmod_mod_trivial [simp]: "(a zmod b) zmod b = a zmod b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_mod_trivial_raw) -apply auto -done - -lemma zmod_zadd_left_eq: "(a$+b) zmod c = ((a zmod c) $+ b) zmod c" -apply (rule trans [symmetric]) -apply (rule zmod_zadd1_eq) -apply (simp (no_asm)) -apply (rule zmod_zadd1_eq [symmetric]) -done - -lemma zmod_zadd_right_eq: "(a$+b) zmod c = (a $+ (b zmod c)) zmod c" -apply (rule trans [symmetric]) -apply (rule zmod_zadd1_eq) -apply (simp (no_asm)) -apply (rule zmod_zadd1_eq [symmetric]) -done - - -lemma zdiv_zadd_self1 [simp]: - "intify(a) \ #0 ==> (a$+b) zdiv a = b zdiv a $+ #1" -by (simp (no_asm_simp) add: zdiv_zadd1_eq) - -lemma zdiv_zadd_self2 [simp]: - "intify(a) \ #0 ==> (b$+a) zdiv a = b zdiv a $+ #1" -by (simp (no_asm_simp) add: zdiv_zadd1_eq) - -lemma zmod_zadd_self1 [simp]: "(a$+b) zmod a = b zmod a" -apply (case_tac "a = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (simp (no_asm_simp) add: zmod_zadd1_eq) -done - -lemma zmod_zadd_self2 [simp]: "(b$+a) zmod a = b zmod a" -apply (case_tac "a = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (simp (no_asm_simp) add: zmod_zadd1_eq) -done - - -subsection\proving a zdiv (b*c) = (a zdiv b) zdiv c\ - -(*The condition c>0 seems necessary. Consider that 7 zdiv ~6 = ~2 but - 7 zdiv 2 zdiv ~3 = 3 zdiv ~3 = ~1. The subcase (a zdiv b) zmod c = 0 seems - to cause particular problems.*) - -(** first, four lemmas to bound the remainder for the cases b<0 and b>0 **) - -lemma zdiv_zmult2_aux1: - "[| #0 $< c; b $< r; r $\ #0 |] ==> b$*c $< b$*(q zmod c) $+ r" -apply (subgoal_tac "b $* (c $- q zmod c) $< r $* #1") -apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) -apply (rule zle_zless_trans) -apply (erule_tac [2] zmult_zless_mono1) -apply (rule zmult_zle_mono2_neg) -apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound) -apply (blast intro: zless_imp_zle dest: zless_zle_trans) -done - -lemma zdiv_zmult2_aux2: - "[| #0 $< c; b $< r; r $\ #0 |] ==> b $* (q zmod c) $+ r $\ #0" -apply (subgoal_tac "b $* (q zmod c) $\ #0") - prefer 2 - apply (simp add: zmult_le_0_iff pos_mod_sign) - apply (blast intro: zless_imp_zle dest: zless_zle_trans) -(*arithmetic*) -apply (drule zadd_zle_mono) -apply assumption -apply (simp add: zadd_commute) -done - -lemma zdiv_zmult2_aux3: - "[| #0 $< c; #0 $\ r; r $< b |] ==> #0 $\ b $* (q zmod c) $+ r" -apply (subgoal_tac "#0 $\ b $* (q zmod c)") - prefer 2 - apply (simp add: int_0_le_mult_iff pos_mod_sign) - apply (blast intro: zless_imp_zle dest: zle_zless_trans) -(*arithmetic*) -apply (drule zadd_zle_mono) -apply assumption -apply (simp add: zadd_commute) -done - -lemma zdiv_zmult2_aux4: - "[| #0 $< c; #0 $\ r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c" -apply (subgoal_tac "r $* #1 $< b $* (c $- q zmod c)") -apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) -apply (rule zless_zle_trans) -apply (erule zmult_zless_mono1) -apply (rule_tac [2] zmult_zle_mono2) -apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound) -apply (blast intro: zless_imp_zle dest: zle_zless_trans) -done - -lemma zdiv_zmult2_lemma: - "[| quorem (, ); a \ int; b \ int; b \ #0; #0 $< c |] - ==> quorem (, )" -apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def - neq_iff_zless int_0_less_mult_iff - zadd_zmult_distrib2 [symmetric] zdiv_zmult2_aux1 zdiv_zmult2_aux2 - zdiv_zmult2_aux3 zdiv_zmult2_aux4) -apply (blast dest: zless_trans)+ -done - -lemma zdiv_zmult2_eq_raw: - "[|#0 $< c; a \ int; b \ int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_div]) -apply (auto simp add: intify_eq_0_iff_zle) -apply (blast dest: zle_zless_trans) -done - -lemma zdiv_zmult2_eq: "#0 $< c ==> a zdiv (b$*c) = (a zdiv b) zdiv c" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zmult2_eq_raw) -apply auto -done - -lemma zmod_zmult2_eq_raw: - "[|#0 $< c; a \ int; b \ int|] - ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_mod]) -apply (auto simp add: intify_eq_0_iff_zle) -apply (blast dest: zle_zless_trans) -done - -lemma zmod_zmult2_eq: - "#0 $< c ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zmult2_eq_raw) -apply auto -done - -subsection\Cancellation of common factors in "zdiv"\ - -lemma zdiv_zmult_zmult1_aux1: - "[| #0 $< b; intify(c) \ #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" -apply (subst zdiv_zmult2_eq) -apply auto -done - -lemma zdiv_zmult_zmult1_aux2: - "[| b $< #0; intify(c) \ #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" -apply (subgoal_tac " (c $* ($-a)) zdiv (c $* ($-b)) = ($-a) zdiv ($-b)") -apply (rule_tac [2] zdiv_zmult_zmult1_aux1) -apply auto -done - -lemma zdiv_zmult_zmult1_raw: - "[|intify(c) \ #0; b \ int|] ==> (c$*a) zdiv (c$*b) = a zdiv b" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (auto simp add: neq_iff_zless [of b] - zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2) -done - -lemma zdiv_zmult_zmult1: "intify(c) \ #0 ==> (c$*a) zdiv (c$*b) = a zdiv b" -apply (cut_tac b = "intify (b)" in zdiv_zmult_zmult1_raw) -apply auto -done - -lemma zdiv_zmult_zmult2: "intify(c) \ #0 ==> (a$*c) zdiv (b$*c) = a zdiv b" -apply (drule zdiv_zmult_zmult1) -apply (auto simp add: zmult_commute) -done - - -subsection\Distribution of factors over "zmod"\ - -lemma zmod_zmult_zmult1_aux1: - "[| #0 $< b; intify(c) \ #0 |] - ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" -apply (subst zmod_zmult2_eq) -apply auto -done - -lemma zmod_zmult_zmult1_aux2: - "[| b $< #0; intify(c) \ #0 |] - ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" -apply (subgoal_tac " (c $* ($-a)) zmod (c $* ($-b)) = c $* (($-a) zmod ($-b))") -apply (rule_tac [2] zmod_zmult_zmult1_aux1) -apply auto -done - -lemma zmod_zmult_zmult1_raw: - "[|b \ int; c \ int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (case_tac "c = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (auto simp add: neq_iff_zless [of b] - zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2) -done - -lemma zmod_zmult_zmult1: "(c$*a) zmod (c$*b) = c $* (a zmod b)" -apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult_zmult1_raw) -apply auto -done - -lemma zmod_zmult_zmult2: "(a$*c) zmod (b$*c) = (a zmod b) $* c" -apply (cut_tac c = "c" in zmod_zmult_zmult1) -apply (auto simp add: zmult_commute) -done - - -(** Quotients of signs **) - -lemma zdiv_neg_pos_less0: "[| a $< #0; #0 $< b |] ==> a zdiv b $< #0" -apply (subgoal_tac "a zdiv b $\ #-1") -apply (erule zle_zless_trans) -apply (simp (no_asm)) -apply (rule zle_trans) -apply (rule_tac a' = "#-1" in zdiv_mono1) -apply (rule zless_add1_iff_zle [THEN iffD1]) -apply (simp (no_asm)) -apply (auto simp add: zdiv_minus1) -done - -lemma zdiv_nonneg_neg_le0: "[| #0 $\ a; b $< #0 |] ==> a zdiv b $\ #0" -apply (drule zdiv_mono1_neg) -apply auto -done - -lemma pos_imp_zdiv_nonneg_iff: "#0 $< b ==> (#0 $\ a zdiv b) \ (#0 $\ a)" -apply auto -apply (drule_tac [2] zdiv_mono1) -apply (auto simp add: neq_iff_zless) -apply (simp (no_asm_use) add: not_zless_iff_zle [THEN iff_sym]) -apply (blast intro: zdiv_neg_pos_less0) -done - -lemma neg_imp_zdiv_nonneg_iff: "b $< #0 ==> (#0 $\ a zdiv b) \ (a $\ #0)" -apply (subst zdiv_zminus_zminus [symmetric]) -apply (rule iff_trans) -apply (rule pos_imp_zdiv_nonneg_iff) -apply auto -done - -(*But not (a zdiv b $\ 0 iff a$\0); consider a=1, b=2 when a zdiv b = 0.*) -lemma pos_imp_zdiv_neg_iff: "#0 $< b ==> (a zdiv b $< #0) \ (a $< #0)" -apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) -apply (erule pos_imp_zdiv_nonneg_iff) -done - -(*Again the law fails for $\: consider a = -1, b = -2 when a zdiv b = 0*) -lemma neg_imp_zdiv_neg_iff: "b $< #0 ==> (a zdiv b $< #0) \ (#0 $< a)" -apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) -apply (erule neg_imp_zdiv_nonneg_iff) -done - -(* - THESE REMAIN TO BE CONVERTED -- but aren't that useful! - - subsection{* Speeding up the division algorithm with shifting *} - - (** computing "zdiv" by shifting **) - - lemma pos_zdiv_mult_2: "#0 $\ a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a" - apply (case_tac "a = #0") - apply (subgoal_tac "#1 $\ a") - apply (arith_tac 2) - apply (subgoal_tac "#1 $< a $* #2") - apply (arith_tac 2) - apply (subgoal_tac "#2$* (#1 $+ b zmod a) $\ #2$*a") - apply (rule_tac [2] zmult_zle_mono2) - apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound) - apply (subst zdiv_zadd1_eq) - apply (simp (no_asm_simp) add: zdiv_zmult_zmult2 zmod_zmult_zmult2 zdiv_pos_pos_trivial) - apply (subst zdiv_pos_pos_trivial) - apply (simp (no_asm_simp) add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ]) - apply (auto simp add: zmod_pos_pos_trivial) - apply (subgoal_tac "#0 $\ b zmod a") - apply (asm_simp_tac (simpset () add: pos_mod_sign) 2) - apply arith - done - - - lemma neg_zdiv_mult_2: "a $\ #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) \ (b$+#1) zdiv a" - apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zdiv (#2 $* ($-a)) \ ($-b-#1) zdiv ($-a)") - apply (rule_tac [2] pos_zdiv_mult_2) - apply (auto simp add: zmult_zminus_right) - apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))") - apply (Simp_tac 2) - apply (asm_full_simp_tac (HOL_ss add: zdiv_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) - done - - - (*Not clear why this must be proved separately; probably integ_of causes - simplification problems*) - lemma lemma: "~ #0 $\ x ==> x $\ #0" - apply auto - done - - lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) = - (if ~b | #0 $\ integ_of w - then integ_of v zdiv (integ_of w) - else (integ_of v $+ #1) zdiv (integ_of w))" - apply (simp_tac (simpset_of @{theory_context Int} add: zadd_assoc integ_of_BIT) - apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zdiv_zmult_zmult1 pos_zdiv_mult_2 lemma neg_zdiv_mult_2) - done - - declare zdiv_integ_of_BIT [simp] - - - (** computing "zmod" by shifting (proofs resemble those for "zdiv") **) - - lemma pos_zmod_mult_2: "#0 $\ a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)" - apply (case_tac "a = #0") - apply (subgoal_tac "#1 $\ a") - apply (arith_tac 2) - apply (subgoal_tac "#1 $< a $* #2") - apply (arith_tac 2) - apply (subgoal_tac "#2$* (#1 $+ b zmod a) $\ #2$*a") - apply (rule_tac [2] zmult_zle_mono2) - apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound) - apply (subst zmod_zadd1_eq) - apply (simp (no_asm_simp) add: zmod_zmult_zmult2 zmod_pos_pos_trivial) - apply (rule zmod_pos_pos_trivial) - apply (simp (no_asm_simp) # add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ]) - apply (auto simp add: zmod_pos_pos_trivial) - apply (subgoal_tac "#0 $\ b zmod a") - apply (asm_simp_tac (simpset () add: pos_mod_sign) 2) - apply arith - done - - - lemma neg_zmod_mult_2: "a $\ #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1" - apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zmod (#2$* ($-a)) = #1 $+ #2$* (($-b-#1) zmod ($-a))") - apply (rule_tac [2] pos_zmod_mult_2) - apply (auto simp add: zmult_zminus_right) - apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))") - apply (Simp_tac 2) - apply (asm_full_simp_tac (HOL_ss add: zmod_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) - apply (dtac (zminus_equation [THEN iffD1, symmetric]) - apply auto - done - - lemma zmod_integ_of_BIT: "integ_of (v BIT b) zmod integ_of (w BIT False) = - (if b then - if #0 $\ integ_of w - then #2 $* (integ_of v zmod integ_of w) $+ #1 - else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1 - else #2 $* (integ_of v zmod integ_of w))" - apply (simp_tac (simpset_of @{theory_context Int} add: zadd_assoc integ_of_BIT) - apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zmod_zmult_zmult1 pos_zmod_mult_2 lemma neg_zmod_mult_2) - done - - declare zmod_integ_of_BIT [simp] -*) - -end - diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Int_ZF.thy --- a/src/ZF/Int_ZF.thy Sun Jun 24 11:41:41 2018 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,921 +0,0 @@ -(* Title: ZF/Int_ZF.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1993 University of Cambridge -*) - -section\The Integers as Equivalence Classes Over Pairs of Natural Numbers\ - -theory Int_ZF imports EquivClass ArithSimp begin - -definition - intrel :: i where - "intrel == {p \ (nat*nat)*(nat*nat). - \x1 y1 x2 y2. p=<,> & x1#+y2 = x2#+y1}" - -definition - int :: i where - "int == (nat*nat)//intrel" - -definition - int_of :: "i=>i" \ \coercion from nat to int\ ("$# _" [80] 80) where - "$# m == intrel `` {}" - -definition - intify :: "i=>i" \ \coercion from ANYTHING to int\ where - "intify(m) == if m \ int then m else $#0" - -definition - raw_zminus :: "i=>i" where - "raw_zminus(z) == \\z. intrel``{}" - -definition - zminus :: "i=>i" ("$- _" [80] 80) where - "$- z == raw_zminus (intify(z))" - -definition - znegative :: "i=>o" where - "znegative(z) == \x y. xnat & \z" - -definition - iszero :: "i=>o" where - "iszero(z) == z = $# 0" - -definition - raw_nat_of :: "i=>i" where - "raw_nat_of(z) == natify (\\z. x#-y)" - -definition - nat_of :: "i=>i" where - "nat_of(z) == raw_nat_of (intify(z))" - -definition - zmagnitude :: "i=>i" where - \ \could be replaced by an absolute value function from int to int?\ - "zmagnitude(z) == - THE m. m\nat & ((~ znegative(z) & z = $# m) | - (znegative(z) & $- z = $# m))" - -definition - raw_zmult :: "[i,i]=>i" where - (*Cannot use UN here or in zadd because of the form of congruent2. - Perhaps a "curried" or even polymorphic congruent predicate would be - better.*) - "raw_zmult(z1,z2) == - \p1\z1. \p2\z2. split(%x1 y1. split(%x2 y2. - intrel``{}, p2), p1)" - -definition - zmult :: "[i,i]=>i" (infixl "$*" 70) where - "z1 $* z2 == raw_zmult (intify(z1),intify(z2))" - -definition - raw_zadd :: "[i,i]=>i" where - "raw_zadd (z1, z2) == - \z1\z1. \z2\z2. let =z1; =z2 - in intrel``{}" - -definition - zadd :: "[i,i]=>i" (infixl "$+" 65) where - "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))" - -definition - zdiff :: "[i,i]=>i" (infixl "$-" 65) where - "z1 $- z2 == z1 $+ zminus(z2)" - -definition - zless :: "[i,i]=>o" (infixl "$<" 50) where - "z1 $< z2 == znegative(z1 $- z2)" - -definition - zle :: "[i,i]=>o" (infixl "$\" 50) where - "z1 $\ z2 == z1 $< z2 | intify(z1)=intify(z2)" - - -declare quotientE [elim!] - -subsection\Proving that @{term intrel} is an equivalence relation\ - -(** Natural deduction for intrel **) - -lemma intrel_iff [simp]: - "<,>: intrel \ - x1\nat & y1\nat & x2\nat & y2\nat & x1#+y2 = x2#+y1" -by (simp add: intrel_def) - -lemma intrelI [intro!]: - "[| x1#+y2 = x2#+y1; x1\nat; y1\nat; x2\nat; y2\nat |] - ==> <,>: intrel" -by (simp add: intrel_def) - -lemma intrelE [elim!]: - "[| p \ intrel; - !!x1 y1 x2 y2. [| p = <,>; x1#+y2 = x2#+y1; - x1\nat; y1\nat; x2\nat; y2\nat |] ==> Q |] - ==> Q" -by (simp add: intrel_def, blast) - -lemma int_trans_lemma: - "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2 |] ==> x1 #+ y3 = x3 #+ y1" -apply (rule sym) -apply (erule add_left_cancel)+ -apply (simp_all (no_asm_simp)) -done - -lemma equiv_intrel: "equiv(nat*nat, intrel)" -apply (simp add: equiv_def refl_def sym_def trans_def) -apply (fast elim!: sym int_trans_lemma) -done - -lemma image_intrel_int: "[| m\nat; n\nat |] ==> intrel `` {} \ int" -by (simp add: int_def) - -declare equiv_intrel [THEN eq_equiv_class_iff, simp] -declare conj_cong [cong] - -lemmas eq_intrelD = eq_equiv_class [OF _ equiv_intrel] - -(** int_of: the injection from nat to int **) - -lemma int_of_type [simp,TC]: "$#m \ int" -by (simp add: int_def quotient_def int_of_def, auto) - -lemma int_of_eq [iff]: "($# m = $# n) \ natify(m)=natify(n)" -by (simp add: int_of_def) - -lemma int_of_inject: "[| $#m = $#n; m\nat; n\nat |] ==> m=n" -by (drule int_of_eq [THEN iffD1], auto) - - -(** intify: coercion from anything to int **) - -lemma intify_in_int [iff,TC]: "intify(x) \ int" -by (simp add: intify_def) - -lemma intify_ident [simp]: "n \ int ==> intify(n) = n" -by (simp add: intify_def) - - -subsection\Collapsing rules: to remove @{term intify} - from arithmetic expressions\ - -lemma intify_idem [simp]: "intify(intify(x)) = intify(x)" -by simp - -lemma int_of_natify [simp]: "$# (natify(m)) = $# m" -by (simp add: int_of_def) - -lemma zminus_intify [simp]: "$- (intify(m)) = $- m" -by (simp add: zminus_def) - -(** Addition **) - -lemma zadd_intify1 [simp]: "intify(x) $+ y = x $+ y" -by (simp add: zadd_def) - -lemma zadd_intify2 [simp]: "x $+ intify(y) = x $+ y" -by (simp add: zadd_def) - -(** Subtraction **) - -lemma zdiff_intify1 [simp]:"intify(x) $- y = x $- y" -by (simp add: zdiff_def) - -lemma zdiff_intify2 [simp]:"x $- intify(y) = x $- y" -by (simp add: zdiff_def) - -(** Multiplication **) - -lemma zmult_intify1 [simp]:"intify(x) $* y = x $* y" -by (simp add: zmult_def) - -lemma zmult_intify2 [simp]:"x $* intify(y) = x $* y" -by (simp add: zmult_def) - -(** Orderings **) - -lemma zless_intify1 [simp]:"intify(x) $< y \ x $< y" -by (simp add: zless_def) - -lemma zless_intify2 [simp]:"x $< intify(y) \ x $< y" -by (simp add: zless_def) - -lemma zle_intify1 [simp]:"intify(x) $\ y \ x $\ y" -by (simp add: zle_def) - -lemma zle_intify2 [simp]:"x $\ intify(y) \ x $\ y" -by (simp add: zle_def) - - -subsection\@{term zminus}: unary negation on @{term int}\ - -lemma zminus_congruent: "(%. intrel``{}) respects intrel" -by (auto simp add: congruent_def add_ac) - -lemma raw_zminus_type: "z \ int ==> raw_zminus(z) \ int" -apply (simp add: int_def raw_zminus_def) -apply (typecheck add: UN_equiv_class_type [OF equiv_intrel zminus_congruent]) -done - -lemma zminus_type [TC,iff]: "$-z \ int" -by (simp add: zminus_def raw_zminus_type) - -lemma raw_zminus_inject: - "[| raw_zminus(z) = raw_zminus(w); z \ int; w \ int |] ==> z=w" -apply (simp add: int_def raw_zminus_def) -apply (erule UN_equiv_class_inject [OF equiv_intrel zminus_congruent], safe) -apply (auto dest: eq_intrelD simp add: add_ac) -done - -lemma zminus_inject_intify [dest!]: "$-z = $-w ==> intify(z) = intify(w)" -apply (simp add: zminus_def) -apply (blast dest!: raw_zminus_inject) -done - -lemma zminus_inject: "[| $-z = $-w; z \ int; w \ int |] ==> z=w" -by auto - -lemma raw_zminus: - "[| x\nat; y\nat |] ==> raw_zminus(intrel``{}) = intrel `` {}" -apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent]) -done - -lemma zminus: - "[| x\nat; y\nat |] - ==> $- (intrel``{}) = intrel `` {}" -by (simp add: zminus_def raw_zminus image_intrel_int) - -lemma raw_zminus_zminus: "z \ int ==> raw_zminus (raw_zminus(z)) = z" -by (auto simp add: int_def raw_zminus) - -lemma zminus_zminus_intify [simp]: "$- ($- z) = intify(z)" -by (simp add: zminus_def raw_zminus_type raw_zminus_zminus) - -lemma zminus_int0 [simp]: "$- ($#0) = $#0" -by (simp add: int_of_def zminus) - -lemma zminus_zminus: "z \ int ==> $- ($- z) = z" -by simp - - -subsection\@{term znegative}: the test for negative integers\ - -lemma znegative: "[| x\nat; y\nat |] ==> znegative(intrel``{}) \ x natify(n)=0" -by (simp add: znegative int_of_def zminus Ord_0_lt_iff [THEN iff_sym]) - - -subsection\@{term nat_of}: Coercion of an Integer to a Natural Number\ - -lemma nat_of_intify [simp]: "nat_of(intify(z)) = nat_of(z)" -by (simp add: nat_of_def) - -lemma nat_of_congruent: "(\x. (\\x,y\. x #- y)(x)) respects intrel" -by (auto simp add: congruent_def split: nat_diff_split) - -lemma raw_nat_of: - "[| x\nat; y\nat |] ==> raw_nat_of(intrel``{}) = x#-y" -by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent]) - -lemma raw_nat_of_int_of: "raw_nat_of($# n) = natify(n)" -by (simp add: int_of_def raw_nat_of) - -lemma nat_of_int_of [simp]: "nat_of($# n) = natify(n)" -by (simp add: raw_nat_of_int_of nat_of_def) - -lemma raw_nat_of_type: "raw_nat_of(z) \ nat" -by (simp add: raw_nat_of_def) - -lemma nat_of_type [iff,TC]: "nat_of(z) \ nat" -by (simp add: nat_of_def raw_nat_of_type) - -subsection\zmagnitude: magnitide of an integer, as a natural number\ - -lemma zmagnitude_int_of [simp]: "zmagnitude($# n) = natify(n)" -by (auto simp add: zmagnitude_def int_of_eq) - -lemma natify_int_of_eq: "natify(x)=n ==> $#x = $# n" -apply (drule sym) -apply (simp (no_asm_simp) add: int_of_eq) -done - -lemma zmagnitude_zminus_int_of [simp]: "zmagnitude($- $# n) = natify(n)" -apply (simp add: zmagnitude_def) -apply (rule the_equality) -apply (auto dest!: not_znegative_imp_zero natify_int_of_eq - iff del: int_of_eq, auto) -done - -lemma zmagnitude_type [iff,TC]: "zmagnitude(z)\nat" -apply (simp add: zmagnitude_def) -apply (rule theI2, auto) -done - -lemma not_zneg_int_of: - "[| z \ int; ~ znegative(z) |] ==> \n\nat. z = $# n" -apply (auto simp add: int_def znegative int_of_def not_lt_iff_le) -apply (rename_tac x y) -apply (rule_tac x="x#-y" in bexI) -apply (auto simp add: add_diff_inverse2) -done - -lemma not_zneg_mag [simp]: - "[| z \ int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z" -by (drule not_zneg_int_of, auto) - -lemma zneg_int_of: - "[| znegative(z); z \ int |] ==> \n\nat. z = $- ($# succ(n))" -by (auto simp add: int_def znegative zminus int_of_def dest!: less_imp_succ_add) - -lemma zneg_mag [simp]: - "[| znegative(z); z \ int |] ==> $# (zmagnitude(z)) = $- z" -by (drule zneg_int_of, auto) - -lemma int_cases: "z \ int ==> \n\nat. z = $# n | z = $- ($# succ(n))" -apply (case_tac "znegative (z) ") -prefer 2 apply (blast dest: not_zneg_mag sym) -apply (blast dest: zneg_int_of) -done - -lemma not_zneg_raw_nat_of: - "[| ~ znegative(z); z \ int |] ==> $# (raw_nat_of(z)) = z" -apply (drule not_zneg_int_of) -apply (auto simp add: raw_nat_of_type raw_nat_of_int_of) -done - -lemma not_zneg_nat_of_intify: - "~ znegative(intify(z)) ==> $# (nat_of(z)) = intify(z)" -by (simp (no_asm_simp) add: nat_of_def not_zneg_raw_nat_of) - -lemma not_zneg_nat_of: "[| ~ znegative(z); z \ int |] ==> $# (nat_of(z)) = z" -apply (simp (no_asm_simp) add: not_zneg_nat_of_intify) -done - -lemma zneg_nat_of [simp]: "znegative(intify(z)) ==> nat_of(z) = 0" -apply (subgoal_tac "intify(z) \ int") -apply (simp add: int_def) -apply (auto simp add: znegative nat_of_def raw_nat_of - split: nat_diff_split) -done - - -subsection\@{term zadd}: addition on int\ - -text\Congruence Property for Addition\ -lemma zadd_congruent2: - "(%z1 z2. let =z1; =z2 - in intrel``{}) - respects2 intrel" -apply (simp add: congruent2_def) -(*Proof via congruent2_commuteI seems longer*) -apply safe -apply (simp (no_asm_simp) add: add_assoc Let_def) -(*The rest should be trivial, but rearranging terms is hard - add_ac does not help rewriting with the assumptions.*) -apply (rule_tac m1 = x1a in add_left_commute [THEN ssubst]) -apply (rule_tac m1 = x2a in add_left_commute [THEN ssubst]) -apply (simp (no_asm_simp) add: add_assoc [symmetric]) -done - -lemma raw_zadd_type: "[| z \ int; w \ int |] ==> raw_zadd(z,w) \ int" -apply (simp add: int_def raw_zadd_def) -apply (rule UN_equiv_class_type2 [OF equiv_intrel zadd_congruent2], assumption+) -apply (simp add: Let_def) -done - -lemma zadd_type [iff,TC]: "z $+ w \ int" -by (simp add: zadd_def raw_zadd_type) - -lemma raw_zadd: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> raw_zadd (intrel``{}, intrel``{}) = - intrel `` {}" -apply (simp add: raw_zadd_def - UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2]) -apply (simp add: Let_def) -done - -lemma zadd: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> (intrel``{}) $+ (intrel``{}) = - intrel `` {}" -by (simp add: zadd_def raw_zadd image_intrel_int) - -lemma raw_zadd_int0: "z \ int ==> raw_zadd ($#0,z) = z" -by (auto simp add: int_def int_of_def raw_zadd) - -lemma zadd_int0_intify [simp]: "$#0 $+ z = intify(z)" -by (simp add: zadd_def raw_zadd_int0) - -lemma zadd_int0: "z \ int ==> $#0 $+ z = z" -by simp - -lemma raw_zminus_zadd_distrib: - "[| z \ int; w \ int |] ==> $- raw_zadd(z,w) = raw_zadd($- z, $- w)" -by (auto simp add: zminus raw_zadd int_def) - -lemma zminus_zadd_distrib [simp]: "$- (z $+ w) = $- z $+ $- w" -by (simp add: zadd_def raw_zminus_zadd_distrib) - -lemma raw_zadd_commute: - "[| z \ int; w \ int |] ==> raw_zadd(z,w) = raw_zadd(w,z)" -by (auto simp add: raw_zadd add_ac int_def) - -lemma zadd_commute: "z $+ w = w $+ z" -by (simp add: zadd_def raw_zadd_commute) - -lemma raw_zadd_assoc: - "[| z1: int; z2: int; z3: int |] - ==> raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))" -by (auto simp add: int_def raw_zadd add_assoc) - -lemma zadd_assoc: "(z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)" -by (simp add: zadd_def raw_zadd_type raw_zadd_assoc) - -(*For AC rewriting*) -lemma zadd_left_commute: "z1$+(z2$+z3) = z2$+(z1$+z3)" -apply (simp add: zadd_assoc [symmetric]) -apply (simp add: zadd_commute) -done - -(*Integer addition is an AC operator*) -lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute - -lemma int_of_add: "$# (m #+ n) = ($#m) $+ ($#n)" -by (simp add: int_of_def zadd) - -lemma int_succ_int_1: "$# succ(m) = $# 1 $+ ($# m)" -by (simp add: int_of_add [symmetric] natify_succ) - -lemma int_of_diff: - "[| m\nat; n \ m |] ==> $# (m #- n) = ($#m) $- ($#n)" -apply (simp add: int_of_def zdiff_def) -apply (frule lt_nat_in_nat) -apply (simp_all add: zadd zminus add_diff_inverse2) -done - -lemma raw_zadd_zminus_inverse: "z \ int ==> raw_zadd (z, $- z) = $#0" -by (auto simp add: int_def int_of_def zminus raw_zadd add_commute) - -lemma zadd_zminus_inverse [simp]: "z $+ ($- z) = $#0" -apply (simp add: zadd_def) -apply (subst zminus_intify [symmetric]) -apply (rule intify_in_int [THEN raw_zadd_zminus_inverse]) -done - -lemma zadd_zminus_inverse2 [simp]: "($- z) $+ z = $#0" -by (simp add: zadd_commute zadd_zminus_inverse) - -lemma zadd_int0_right_intify [simp]: "z $+ $#0 = intify(z)" -by (rule trans [OF zadd_commute zadd_int0_intify]) - -lemma zadd_int0_right: "z \ int ==> z $+ $#0 = z" -by simp - - -subsection\@{term zmult}: Integer Multiplication\ - -text\Congruence property for multiplication\ -lemma zmult_congruent2: - "(%p1 p2. split(%x1 y1. split(%x2 y2. - intrel``{}, p2), p1)) - respects2 intrel" -apply (rule equiv_intrel [THEN congruent2_commuteI], auto) -(*Proof that zmult is congruent in one argument*) -apply (rename_tac x y) -apply (frule_tac t = "%u. x#*u" in sym [THEN subst_context]) -apply (drule_tac t = "%u. y#*u" in subst_context) -apply (erule add_left_cancel)+ -apply (simp_all add: add_mult_distrib_left) -done - - -lemma raw_zmult_type: "[| z \ int; w \ int |] ==> raw_zmult(z,w) \ int" -apply (simp add: int_def raw_zmult_def) -apply (rule UN_equiv_class_type2 [OF equiv_intrel zmult_congruent2], assumption+) -apply (simp add: Let_def) -done - -lemma zmult_type [iff,TC]: "z $* w \ int" -by (simp add: zmult_def raw_zmult_type) - -lemma raw_zmult: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> raw_zmult(intrel``{}, intrel``{}) = - intrel `` {}" -by (simp add: raw_zmult_def - UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2]) - -lemma zmult: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> (intrel``{}) $* (intrel``{}) = - intrel `` {}" -by (simp add: zmult_def raw_zmult image_intrel_int) - -lemma raw_zmult_int0: "z \ int ==> raw_zmult ($#0,z) = $#0" -by (auto simp add: int_def int_of_def raw_zmult) - -lemma zmult_int0 [simp]: "$#0 $* z = $#0" -by (simp add: zmult_def raw_zmult_int0) - -lemma raw_zmult_int1: "z \ int ==> raw_zmult ($#1,z) = z" -by (auto simp add: int_def int_of_def raw_zmult) - -lemma zmult_int1_intify [simp]: "$#1 $* z = intify(z)" -by (simp add: zmult_def raw_zmult_int1) - -lemma zmult_int1: "z \ int ==> $#1 $* z = z" -by simp - -lemma raw_zmult_commute: - "[| z \ int; w \ int |] ==> raw_zmult(z,w) = raw_zmult(w,z)" -by (auto simp add: int_def raw_zmult add_ac mult_ac) - -lemma zmult_commute: "z $* w = w $* z" -by (simp add: zmult_def raw_zmult_commute) - -lemma raw_zmult_zminus: - "[| z \ int; w \ int |] ==> raw_zmult($- z, w) = $- raw_zmult(z, w)" -by (auto simp add: int_def zminus raw_zmult add_ac) - -lemma zmult_zminus [simp]: "($- z) $* w = $- (z $* w)" -apply (simp add: zmult_def raw_zmult_zminus) -apply (subst zminus_intify [symmetric], rule raw_zmult_zminus, auto) -done - -lemma zmult_zminus_right [simp]: "w $* ($- z) = $- (w $* z)" -by (simp add: zmult_commute [of w]) - -lemma raw_zmult_assoc: - "[| z1: int; z2: int; z3: int |] - ==> raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))" -by (auto simp add: int_def raw_zmult add_mult_distrib_left add_ac mult_ac) - -lemma zmult_assoc: "(z1 $* z2) $* z3 = z1 $* (z2 $* z3)" -by (simp add: zmult_def raw_zmult_type raw_zmult_assoc) - -(*For AC rewriting*) -lemma zmult_left_commute: "z1$*(z2$*z3) = z2$*(z1$*z3)" -apply (simp add: zmult_assoc [symmetric]) -apply (simp add: zmult_commute) -done - -(*Integer multiplication is an AC operator*) -lemmas zmult_ac = zmult_assoc zmult_commute zmult_left_commute - -lemma raw_zadd_zmult_distrib: - "[| z1: int; z2: int; w \ int |] - ==> raw_zmult(raw_zadd(z1,z2), w) = - raw_zadd (raw_zmult(z1,w), raw_zmult(z2,w))" -by (auto simp add: int_def raw_zadd raw_zmult add_mult_distrib_left add_ac mult_ac) - -lemma zadd_zmult_distrib: "(z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)" -by (simp add: zmult_def zadd_def raw_zadd_type raw_zmult_type - raw_zadd_zmult_distrib) - -lemma zadd_zmult_distrib2: "w $* (z1 $+ z2) = (w $* z1) $+ (w $* z2)" -by (simp add: zmult_commute [of w] zadd_zmult_distrib) - -lemmas int_typechecks = - int_of_type zminus_type zmagnitude_type zadd_type zmult_type - - -(*** Subtraction laws ***) - -lemma zdiff_type [iff,TC]: "z $- w \ int" -by (simp add: zdiff_def) - -lemma zminus_zdiff_eq [simp]: "$- (z $- y) = y $- z" -by (simp add: zdiff_def zadd_commute) - -lemma zdiff_zmult_distrib: "(z1 $- z2) $* w = (z1 $* w) $- (z2 $* w)" -apply (simp add: zdiff_def) -apply (subst zadd_zmult_distrib) -apply (simp add: zmult_zminus) -done - -lemma zdiff_zmult_distrib2: "w $* (z1 $- z2) = (w $* z1) $- (w $* z2)" -by (simp add: zmult_commute [of w] zdiff_zmult_distrib) - -lemma zadd_zdiff_eq: "x $+ (y $- z) = (x $+ y) $- z" -by (simp add: zdiff_def zadd_ac) - -lemma zdiff_zadd_eq: "(x $- y) $+ z = (x $+ z) $- y" -by (simp add: zdiff_def zadd_ac) - - -subsection\The "Less Than" Relation\ - -(*"Less than" is a linear ordering*) -lemma zless_linear_lemma: - "[| z \ int; w \ int |] ==> z$ int; y \ int |] ==> (x \ y) \ (x $< y | y $< x)" -by (cut_tac z = x and w = y in zless_linear, auto) - -lemma zless_imp_intify_neq: "w $< z ==> intify(w) \ intify(z)" -apply auto -apply (subgoal_tac "~ (intify (w) $< intify (z))") -apply (erule_tac [2] ssubst) -apply (simp (no_asm_use)) -apply auto -done - -(*This lemma allows direct proofs of other <-properties*) -lemma zless_imp_succ_zadd_lemma: - "[| w $< z; w \ int; z \ int |] ==> (\n\nat. z = w $+ $#(succ(n)))" -apply (simp add: zless_def znegative_def zdiff_def int_def) -apply (auto dest!: less_imp_succ_add simp add: zadd zminus int_of_def) -apply (rule_tac x = k in bexI) -apply (erule_tac i="succ (v)" for v in add_left_cancel, auto) -done - -lemma zless_imp_succ_zadd: - "w $< z ==> (\n\nat. w $+ $#(succ(n)) = intify(z))" -apply (subgoal_tac "intify (w) $< intify (z) ") -apply (drule_tac w = "intify (w) " in zless_imp_succ_zadd_lemma) -apply auto -done - -lemma zless_succ_zadd_lemma: - "w \ int ==> w $< w $+ $# succ(n)" -apply (simp add: zless_def znegative_def zdiff_def int_def) -apply (auto simp add: zadd zminus int_of_def image_iff) -apply (rule_tac x = 0 in exI, auto) -done - -lemma zless_succ_zadd: "w $< w $+ $# succ(n)" -by (cut_tac intify_in_int [THEN zless_succ_zadd_lemma], auto) - -lemma zless_iff_succ_zadd: - "w $< z \ (\n\nat. w $+ $#(succ(n)) = intify(z))" -apply (rule iffI) -apply (erule zless_imp_succ_zadd, auto) -apply (rename_tac "n") -apply (cut_tac w = w and n = n in zless_succ_zadd, auto) -done - -lemma zless_int_of [simp]: "[| m\nat; n\nat |] ==> ($#m $< $#n) \ (m int; y \ int; z \ int |] ==> x $< z" -apply (simp add: zless_def znegative_def zdiff_def int_def) -apply (auto simp add: zadd zminus image_iff) -apply (rename_tac x1 x2 y1 y2) -apply (rule_tac x = "x1#+x2" in exI) -apply (rule_tac x = "y1#+y2" in exI) -apply (auto simp add: add_lt_mono) -apply (rule sym) -apply hypsubst_thin -apply (erule add_left_cancel)+ -apply auto -done - -lemma zless_trans [trans]: "[| x $< y; y $< z |] ==> x $< z" -apply (subgoal_tac "intify (x) $< intify (z) ") -apply (rule_tac [2] y = "intify (y) " in zless_trans_lemma) -apply auto -done - -lemma zless_not_sym: "z $< w ==> ~ (w $< z)" -by (blast dest: zless_trans) - -(* [| z $< w; ~ P ==> w $< z |] ==> P *) -lemmas zless_asym = zless_not_sym [THEN swap] - -lemma zless_imp_zle: "z $< w ==> z $\ w" -by (simp add: zle_def) - -lemma zle_linear: "z $\ w | w $\ z" -apply (simp add: zle_def) -apply (cut_tac zless_linear, blast) -done - - -subsection\Less Than or Equals\ - -lemma zle_refl: "z $\ z" -by (simp add: zle_def) - -lemma zle_eq_refl: "x=y ==> x $\ y" -by (simp add: zle_refl) - -lemma zle_anti_sym_intify: "[| x $\ y; y $\ x |] ==> intify(x) = intify(y)" -apply (simp add: zle_def, auto) -apply (blast dest: zless_trans) -done - -lemma zle_anti_sym: "[| x $\ y; y $\ x; x \ int; y \ int |] ==> x=y" -by (drule zle_anti_sym_intify, auto) - -lemma zle_trans_lemma: - "[| x \ int; y \ int; z \ int; x $\ y; y $\ z |] ==> x $\ z" -apply (simp add: zle_def, auto) -apply (blast intro: zless_trans) -done - -lemma zle_trans [trans]: "[| x $\ y; y $\ z |] ==> x $\ z" -apply (subgoal_tac "intify (x) $\ intify (z) ") -apply (rule_tac [2] y = "intify (y) " in zle_trans_lemma) -apply auto -done - -lemma zle_zless_trans [trans]: "[| i $\ j; j $< k |] ==> i $< k" -apply (auto simp add: zle_def) -apply (blast intro: zless_trans) -apply (simp add: zless_def zdiff_def zadd_def) -done - -lemma zless_zle_trans [trans]: "[| i $< j; j $\ k |] ==> i $< k" -apply (auto simp add: zle_def) -apply (blast intro: zless_trans) -apply (simp add: zless_def zdiff_def zminus_def) -done - -lemma not_zless_iff_zle: "~ (z $< w) \ (w $\ z)" -apply (cut_tac z = z and w = w in zless_linear) -apply (auto dest: zless_trans simp add: zle_def) -apply (auto dest!: zless_imp_intify_neq) -done - -lemma not_zle_iff_zless: "~ (z $\ w) \ (w $< z)" -by (simp add: not_zless_iff_zle [THEN iff_sym]) - - -subsection\More subtraction laws (for \zcompare_rls\)\ - -lemma zdiff_zdiff_eq: "(x $- y) $- z = x $- (y $+ z)" -by (simp add: zdiff_def zadd_ac) - -lemma zdiff_zdiff_eq2: "x $- (y $- z) = (x $+ z) $- y" -by (simp add: zdiff_def zadd_ac) - -lemma zdiff_zless_iff: "(x$-y $< z) \ (x $< z $+ y)" -by (simp add: zless_def zdiff_def zadd_ac) - -lemma zless_zdiff_iff: "(x $< z$-y) \ (x $+ y $< z)" -by (simp add: zless_def zdiff_def zadd_ac) - -lemma zdiff_eq_iff: "[| x \ int; z \ int |] ==> (x$-y = z) \ (x = z $+ y)" -by (auto simp add: zdiff_def zadd_assoc) - -lemma eq_zdiff_iff: "[| x \ int; z \ int |] ==> (x = z$-y) \ (x $+ y = z)" -by (auto simp add: zdiff_def zadd_assoc) - -lemma zdiff_zle_iff_lemma: - "[| x \ int; z \ int |] ==> (x$-y $\ z) \ (x $\ z $+ y)" -by (auto simp add: zle_def zdiff_eq_iff zdiff_zless_iff) - -lemma zdiff_zle_iff: "(x$-y $\ z) \ (x $\ z $+ y)" -by (cut_tac zdiff_zle_iff_lemma [OF intify_in_int intify_in_int], simp) - -lemma zle_zdiff_iff_lemma: - "[| x \ int; z \ int |] ==>(x $\ z$-y) \ (x $+ y $\ z)" -apply (auto simp add: zle_def zdiff_eq_iff zless_zdiff_iff) -apply (auto simp add: zdiff_def zadd_assoc) -done - -lemma zle_zdiff_iff: "(x $\ z$-y) \ (x $+ y $\ z)" -by (cut_tac zle_zdiff_iff_lemma [ OF intify_in_int intify_in_int], simp) - -text\This list of rewrites simplifies (in)equalities by bringing subtractions - to the top and then moving negative terms to the other side. - Use with \zadd_ac\\ -lemmas zcompare_rls = - zdiff_def [symmetric] - zadd_zdiff_eq zdiff_zadd_eq zdiff_zdiff_eq zdiff_zdiff_eq2 - zdiff_zless_iff zless_zdiff_iff zdiff_zle_iff zle_zdiff_iff - zdiff_eq_iff eq_zdiff_iff - - -subsection\Monotonicity and Cancellation Results for Instantiation - of the CancelNumerals Simprocs\ - -lemma zadd_left_cancel: - "[| w \ int; w': int |] ==> (z $+ w' = z $+ w) \ (w' = w)" -apply safe -apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) -apply (simp add: zadd_ac) -done - -lemma zadd_left_cancel_intify [simp]: - "(z $+ w' = z $+ w) \ intify(w') = intify(w)" -apply (rule iff_trans) -apply (rule_tac [2] zadd_left_cancel, auto) -done - -lemma zadd_right_cancel: - "[| w \ int; w': int |] ==> (w' $+ z = w $+ z) \ (w' = w)" -apply safe -apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) -apply (simp add: zadd_ac) -done - -lemma zadd_right_cancel_intify [simp]: - "(w' $+ z = w $+ z) \ intify(w') = intify(w)" -apply (rule iff_trans) -apply (rule_tac [2] zadd_right_cancel, auto) -done - -lemma zadd_right_cancel_zless [simp]: "(w' $+ z $< w $+ z) \ (w' $< w)" -by (simp add: zdiff_zless_iff [THEN iff_sym] zdiff_def zadd_assoc) - -lemma zadd_left_cancel_zless [simp]: "(z $+ w' $< z $+ w) \ (w' $< w)" -by (simp add: zadd_commute [of z] zadd_right_cancel_zless) - -lemma zadd_right_cancel_zle [simp]: "(w' $+ z $\ w $+ z) \ w' $\ w" -by (simp add: zle_def) - -lemma zadd_left_cancel_zle [simp]: "(z $+ w' $\ z $+ w) \ w' $\ w" -by (simp add: zadd_commute [of z] zadd_right_cancel_zle) - - -(*"v $\ w ==> v$+z $\ w$+z"*) -lemmas zadd_zless_mono1 = zadd_right_cancel_zless [THEN iffD2] - -(*"v $\ w ==> z$+v $\ z$+w"*) -lemmas zadd_zless_mono2 = zadd_left_cancel_zless [THEN iffD2] - -(*"v $\ w ==> v$+z $\ w$+z"*) -lemmas zadd_zle_mono1 = zadd_right_cancel_zle [THEN iffD2] - -(*"v $\ w ==> z$+v $\ z$+w"*) -lemmas zadd_zle_mono2 = zadd_left_cancel_zle [THEN iffD2] - -lemma zadd_zle_mono: "[| w' $\ w; z' $\ z |] ==> w' $+ z' $\ w $+ z" -by (erule zadd_zle_mono1 [THEN zle_trans], simp) - -lemma zadd_zless_mono: "[| w' $< w; z' $\ z |] ==> w' $+ z' $< w $+ z" -by (erule zadd_zless_mono1 [THEN zless_zle_trans], simp) - - -subsection\Comparison laws\ - -lemma zminus_zless_zminus [simp]: "($- x $< $- y) \ (y $< x)" -by (simp add: zless_def zdiff_def zadd_ac) - -lemma zminus_zle_zminus [simp]: "($- x $\ $- y) \ (y $\ x)" -by (simp add: not_zless_iff_zle [THEN iff_sym]) - -subsubsection\More inequality lemmas\ - -lemma equation_zminus: "[| x \ int; y \ int |] ==> (x = $- y) \ (y = $- x)" -by auto - -lemma zminus_equation: "[| x \ int; y \ int |] ==> ($- x = y) \ ($- y = x)" -by auto - -lemma equation_zminus_intify: "(intify(x) = $- y) \ (intify(y) = $- x)" -apply (cut_tac x = "intify (x) " and y = "intify (y) " in equation_zminus) -apply auto -done - -lemma zminus_equation_intify: "($- x = intify(y)) \ ($- y = intify(x))" -apply (cut_tac x = "intify (x) " and y = "intify (y) " in zminus_equation) -apply auto -done - - -subsubsection\The next several equations are permutative: watch out!\ - -lemma zless_zminus: "(x $< $- y) \ (y $< $- x)" -by (simp add: zless_def zdiff_def zadd_ac) - -lemma zminus_zless: "($- x $< y) \ ($- y $< x)" -by (simp add: zless_def zdiff_def zadd_ac) - -lemma zle_zminus: "(x $\ $- y) \ (y $\ $- x)" -by (simp add: not_zless_iff_zle [THEN iff_sym] zminus_zless) - -lemma zminus_zle: "($- x $\ y) \ ($- y $\ x)" -by (simp add: not_zless_iff_zle [THEN iff_sym] zless_zminus) - -end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/List.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/List.thy Sun Jun 24 15:57:48 2018 +0200 @@ -0,0 +1,1249 @@ +(* Title: ZF/List.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge +*) + +section\Lists in Zermelo-Fraenkel Set Theory\ + +theory List imports Datatype ArithSimp begin + +consts + list :: "i=>i" + +datatype + "list(A)" = Nil | Cons ("a \ A", "l \ list(A)") + + +syntax + "_Nil" :: i ("[]") + "_List" :: "is => i" ("[(_)]") + +translations + "[x, xs]" == "CONST Cons(x, [xs])" + "[x]" == "CONST Cons(x, [])" + "[]" == "CONST Nil" + + +consts + length :: "i=>i" + hd :: "i=>i" + tl :: "i=>i" + +primrec + "length([]) = 0" + "length(Cons(a,l)) = succ(length(l))" + +primrec + "hd([]) = 0" + "hd(Cons(a,l)) = a" + +primrec + "tl([]) = []" + "tl(Cons(a,l)) = l" + + +consts + map :: "[i=>i, i] => i" + set_of_list :: "i=>i" + app :: "[i,i]=>i" (infixr "@" 60) + +(*map is a binding operator -- it applies to meta-level functions, not +object-level functions. This simplifies the final form of term_rec_conv, +although complicating its derivation.*) +primrec + "map(f,[]) = []" + "map(f,Cons(a,l)) = Cons(f(a), map(f,l))" + +primrec + "set_of_list([]) = 0" + "set_of_list(Cons(a,l)) = cons(a, set_of_list(l))" + +primrec + app_Nil: "[] @ ys = ys" + app_Cons: "(Cons(a,l)) @ ys = Cons(a, l @ ys)" + + +consts + rev :: "i=>i" + flat :: "i=>i" + list_add :: "i=>i" + +primrec + "rev([]) = []" + "rev(Cons(a,l)) = rev(l) @ [a]" + +primrec + "flat([]) = []" + "flat(Cons(l,ls)) = l @ flat(ls)" + +primrec + "list_add([]) = 0" + "list_add(Cons(a,l)) = a #+ list_add(l)" + +consts + drop :: "[i,i]=>i" + +primrec + drop_0: "drop(0,l) = l" + drop_succ: "drop(succ(i), l) = tl (drop(i,l))" + + +(*** Thanks to Sidi Ehmety for the following ***) + +definition +(* Function `take' returns the first n elements of a list *) + take :: "[i,i]=>i" where + "take(n, as) == list_rec(\n\nat. [], + %a l r. \n\nat. nat_case([], %m. Cons(a, r`m), n), as)`n" + +definition + nth :: "[i, i]=>i" where + \ \returns the (n+1)th element of a list, or 0 if the + list is too short.\ + "nth(n, as) == list_rec(\n\nat. 0, + %a l r. \n\nat. nat_case(a, %m. r`m, n), as) ` n" + +definition + list_update :: "[i, i, i]=>i" where + "list_update(xs, i, v) == list_rec(\n\nat. Nil, + %u us vs. \n\nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i" + +consts + filter :: "[i=>o, i] => i" + upt :: "[i, i] =>i" + +primrec + "filter(P, Nil) = Nil" + "filter(P, Cons(x, xs)) = + (if P(x) then Cons(x, filter(P, xs)) else filter(P, xs))" + +primrec + "upt(i, 0) = Nil" + "upt(i, succ(j)) = (if i \ j then upt(i, j)@[j] else Nil)" + +definition + min :: "[i,i] =>i" where + "min(x, y) == (if x \ y then x else y)" + +definition + max :: "[i, i] =>i" where + "max(x, y) == (if x \ y then y else x)" + +(*** Aspects of the datatype definition ***) + +declare list.intros [simp,TC] + +(*An elimination rule, for type-checking*) +inductive_cases ConsE: "Cons(a,l) \ list(A)" + +lemma Cons_type_iff [simp]: "Cons(a,l) \ list(A) \ a \ A & l \ list(A)" +by (blast elim: ConsE) + +(*Proving freeness results*) +lemma Cons_iff: "Cons(a,l)=Cons(a',l') \ a=a' & l=l'" +by auto + +lemma Nil_Cons_iff: "~ Nil=Cons(a,l)" +by auto + +lemma list_unfold: "list(A) = {0} + (A * list(A))" +by (blast intro!: list.intros [unfolded list.con_defs] + elim: list.cases [unfolded list.con_defs]) + + +(** Lemmas to justify using "list" in other recursive type definitions **) + +lemma list_mono: "A<=B ==> list(A) \ list(B)" +apply (unfold list.defs ) +apply (rule lfp_mono) +apply (simp_all add: list.bnd_mono) +apply (assumption | rule univ_mono basic_monos)+ +done + +(*There is a similar proof by list induction.*) +lemma list_univ: "list(univ(A)) \ univ(A)" +apply (unfold list.defs list.con_defs) +apply (rule lfp_lowerbound) +apply (rule_tac [2] A_subset_univ [THEN univ_mono]) +apply (blast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ) +done + +(*These two theorems justify datatypes involving list(nat), list(A), ...*) +lemmas list_subset_univ = subset_trans [OF list_mono list_univ] + +lemma list_into_univ: "[| l \ list(A); A \ univ(B) |] ==> l \ univ(B)" +by (blast intro: list_subset_univ [THEN subsetD]) + +lemma list_case_type: + "[| l \ list(A); + c \ C(Nil); + !!x y. [| x \ A; y \ list(A) |] ==> h(x,y): C(Cons(x,y)) + |] ==> list_case(c,h,l) \ C(l)" +by (erule list.induct, auto) + +lemma list_0_triv: "list(0) = {Nil}" +apply (rule equalityI, auto) +apply (induct_tac x, auto) +done + + +(*** List functions ***) + +lemma tl_type: "l \ list(A) ==> tl(l) \ list(A)" +apply (induct_tac "l") +apply (simp_all (no_asm_simp) add: list.intros) +done + +(** drop **) + +lemma drop_Nil [simp]: "i \ nat ==> drop(i, Nil) = Nil" +apply (induct_tac "i") +apply (simp_all (no_asm_simp)) +done + +lemma drop_succ_Cons [simp]: "i \ nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)" +apply (rule sym) +apply (induct_tac "i") +apply (simp (no_asm)) +apply (simp (no_asm_simp)) +done + +lemma drop_type [simp,TC]: "[| i \ nat; l \ list(A) |] ==> drop(i,l) \ list(A)" +apply (induct_tac "i") +apply (simp_all (no_asm_simp) add: tl_type) +done + +declare drop_succ [simp del] + + +(** Type checking -- proved by induction, as usual **) + +lemma list_rec_type [TC]: + "[| l \ list(A); + c \ C(Nil); + !!x y r. [| x \ A; y \ list(A); r \ C(y) |] ==> h(x,y,r): C(Cons(x,y)) + |] ==> list_rec(c,h,l) \ C(l)" +by (induct_tac "l", auto) + +(** map **) + +lemma map_type [TC]: + "[| l \ list(A); !!x. x \ A ==> h(x): B |] ==> map(h,l) \ list(B)" +apply (simp add: map_list_def) +apply (typecheck add: list.intros list_rec_type, blast) +done + +lemma map_type2 [TC]: "l \ list(A) ==> map(h,l) \ list({h(u). u \ A})" +apply (erule map_type) +apply (erule RepFunI) +done + +(** length **) + +lemma length_type [TC]: "l \ list(A) ==> length(l) \ nat" +by (simp add: length_list_def) + +lemma lt_length_in_nat: + "[|x < length(xs); xs \ list(A)|] ==> x \ nat" +by (frule lt_nat_in_nat, typecheck) + +(** app **) + +lemma app_type [TC]: "[| xs: list(A); ys: list(A) |] ==> xs@ys \ list(A)" +by (simp add: app_list_def) + +(** rev **) + +lemma rev_type [TC]: "xs: list(A) ==> rev(xs) \ list(A)" +by (simp add: rev_list_def) + + +(** flat **) + +lemma flat_type [TC]: "ls: list(list(A)) ==> flat(ls) \ list(A)" +by (simp add: flat_list_def) + + +(** set_of_list **) + +lemma set_of_list_type [TC]: "l \ list(A) ==> set_of_list(l) \ Pow(A)" +apply (unfold set_of_list_list_def) +apply (erule list_rec_type, auto) +done + +lemma set_of_list_append: + "xs: list(A) ==> set_of_list (xs@ys) = set_of_list(xs) \ set_of_list(ys)" +apply (erule list.induct) +apply (simp_all (no_asm_simp) add: Un_cons) +done + + +(** list_add **) + +lemma list_add_type [TC]: "xs: list(nat) ==> list_add(xs) \ nat" +by (simp add: list_add_list_def) + + +(*** theorems about map ***) + +lemma map_ident [simp]: "l \ list(A) ==> map(%u. u, l) = l" +apply (induct_tac "l") +apply (simp_all (no_asm_simp)) +done + +lemma map_compose: "l \ list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)" +apply (induct_tac "l") +apply (simp_all (no_asm_simp)) +done + +lemma map_app_distrib: "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)" +apply (induct_tac "xs") +apply (simp_all (no_asm_simp)) +done + +lemma map_flat: "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))" +apply (induct_tac "ls") +apply (simp_all (no_asm_simp) add: map_app_distrib) +done + +lemma list_rec_map: + "l \ list(A) ==> + list_rec(c, d, map(h,l)) = + list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)" +apply (induct_tac "l") +apply (simp_all (no_asm_simp)) +done + +(** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **) + +(* @{term"c \ list(Collect(B,P)) ==> c \ list"} *) +lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD] + +lemma map_list_Collect: "l \ list({x \ A. h(x)=j(x)}) ==> map(h,l) = map(j,l)" +apply (induct_tac "l") +apply (simp_all (no_asm_simp)) +done + +(*** theorems about length ***) + +lemma length_map [simp]: "xs: list(A) ==> length(map(h,xs)) = length(xs)" +by (induct_tac "xs", simp_all) + +lemma length_app [simp]: + "[| xs: list(A); ys: list(A) |] + ==> length(xs@ys) = length(xs) #+ length(ys)" +by (induct_tac "xs", simp_all) + +lemma length_rev [simp]: "xs: list(A) ==> length(rev(xs)) = length(xs)" +apply (induct_tac "xs") +apply (simp_all (no_asm_simp) add: length_app) +done + +lemma length_flat: + "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))" +apply (induct_tac "ls") +apply (simp_all (no_asm_simp) add: length_app) +done + +(** Length and drop **) + +(*Lemma for the inductive step of drop_length*) +lemma drop_length_Cons [rule_format]: + "xs: list(A) ==> + \x. \z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)" +by (erule list.induct, simp_all) + +lemma drop_length [rule_format]: + "l \ list(A) ==> \i \ length(l). (\z zs. drop(i,l) = Cons(z,zs))" +apply (erule list.induct, simp_all, safe) +apply (erule drop_length_Cons) +apply (rule natE) +apply (erule Ord_trans [OF asm_rl length_type Ord_nat], assumption, simp_all) +apply (blast intro: succ_in_naturalD length_type) +done + + +(*** theorems about app ***) + +lemma app_right_Nil [simp]: "xs: list(A) ==> xs@Nil=xs" +by (erule list.induct, simp_all) + +lemma app_assoc: "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)" +by (induct_tac "xs", simp_all) + +lemma flat_app_distrib: "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)" +apply (induct_tac "ls") +apply (simp_all (no_asm_simp) add: app_assoc) +done + +(*** theorems about rev ***) + +lemma rev_map_distrib: "l \ list(A) ==> rev(map(h,l)) = map(h,rev(l))" +apply (induct_tac "l") +apply (simp_all (no_asm_simp) add: map_app_distrib) +done + +(*Simplifier needs the premises as assumptions because rewriting will not + instantiate the variable ?A in the rules' typing conditions; note that + rev_type does not instantiate ?A. Only the premises do. +*) +lemma rev_app_distrib: + "[| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)" +apply (erule list.induct) +apply (simp_all add: app_assoc) +done + +lemma rev_rev_ident [simp]: "l \ list(A) ==> rev(rev(l))=l" +apply (induct_tac "l") +apply (simp_all (no_asm_simp) add: rev_app_distrib) +done + +lemma rev_flat: "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))" +apply (induct_tac "ls") +apply (simp_all add: map_app_distrib flat_app_distrib rev_app_distrib) +done + + +(*** theorems about list_add ***) + +lemma list_add_app: + "[| xs: list(nat); ys: list(nat) |] + ==> list_add(xs@ys) = list_add(ys) #+ list_add(xs)" +apply (induct_tac "xs", simp_all) +done + +lemma list_add_rev: "l \ list(nat) ==> list_add(rev(l)) = list_add(l)" +apply (induct_tac "l") +apply (simp_all (no_asm_simp) add: list_add_app) +done + +lemma list_add_flat: + "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))" +apply (induct_tac "ls") +apply (simp_all (no_asm_simp) add: list_add_app) +done + +(** New induction rules **) + +lemma list_append_induct [case_names Nil snoc, consumes 1]: + "[| l \ list(A); + P(Nil); + !!x y. [| x \ A; y \ list(A); P(y) |] ==> P(y @ [x]) + |] ==> P(l)" +apply (subgoal_tac "P(rev(rev(l)))", simp) +apply (erule rev_type [THEN list.induct], simp_all) +done + +lemma list_complete_induct_lemma [rule_format]: + assumes ih: + "\l. [| l \ list(A); + \l' \ list(A). length(l') < length(l) \ P(l')|] + ==> P(l)" + shows "n \ nat ==> \l \ list(A). length(l) < n \ P(l)" +apply (induct_tac n, simp) +apply (blast intro: ih elim!: leE) +done + +theorem list_complete_induct: + "[| l \ list(A); + \l. [| l \ list(A); + \l' \ list(A). length(l') < length(l) \ P(l')|] + ==> P(l) + |] ==> P(l)" +apply (rule list_complete_induct_lemma [of A]) + prefer 4 apply (rule le_refl, simp) + apply blast + apply simp +apply assumption +done + + +(*** Thanks to Sidi Ehmety for these results about min, take, etc. ***) + +(** min FIXME: replace by Int! **) +(* Min theorems are also true for i, j ordinals *) +lemma min_sym: "[| i \ nat; j \ nat |] ==> min(i,j)=min(j,i)" +apply (unfold min_def) +apply (auto dest!: not_lt_imp_le dest: lt_not_sym intro: le_anti_sym) +done + +lemma min_type [simp,TC]: "[| i \ nat; j \ nat |] ==> min(i,j):nat" +by (unfold min_def, auto) + +lemma min_0 [simp]: "i \ nat ==> min(0,i) = 0" +apply (unfold min_def) +apply (auto dest: not_lt_imp_le) +done + +lemma min_02 [simp]: "i \ nat ==> min(i, 0) = 0" +apply (unfold min_def) +apply (auto dest: not_lt_imp_le) +done + +lemma lt_min_iff: "[| i \ nat; j \ nat; k \ nat |] ==> i i nat; j \ nat |] ==> min(succ(i), succ(j))= succ(min(i, j))" +apply (unfold min_def, auto) +done + +(*** more theorems about lists ***) + +(** filter **) + +lemma filter_append [simp]: + "xs:list(A) ==> filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)" +by (induct_tac "xs", auto) + +lemma filter_type [simp,TC]: "xs:list(A) ==> filter(P, xs):list(A)" +by (induct_tac "xs", auto) + +lemma length_filter: "xs:list(A) ==> length(filter(P, xs)) \ length(xs)" +apply (induct_tac "xs", auto) +apply (rule_tac j = "length (l) " in le_trans) +apply (auto simp add: le_iff) +done + +lemma filter_is_subset: "xs:list(A) ==> set_of_list(filter(P,xs)) \ set_of_list(xs)" +by (induct_tac "xs", auto) + +lemma filter_False [simp]: "xs:list(A) ==> filter(%p. False, xs) = Nil" +by (induct_tac "xs", auto) + +lemma filter_True [simp]: "xs:list(A) ==> filter(%p. True, xs) = xs" +by (induct_tac "xs", auto) + +(** length **) + +lemma length_is_0_iff [simp]: "xs:list(A) ==> length(xs)=0 \ xs=Nil" +by (erule list.induct, auto) + +lemma length_is_0_iff2 [simp]: "xs:list(A) ==> 0 = length(xs) \ xs=Nil" +by (erule list.induct, auto) + +lemma length_tl [simp]: "xs:list(A) ==> length(tl(xs)) = length(xs) #- 1" +by (erule list.induct, auto) + +lemma length_greater_0_iff: "xs:list(A) ==> 0 xs \ Nil" +by (erule list.induct, auto) + +lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) \ (\y ys. xs=Cons(y, ys) & length(ys)=n)" +by (erule list.induct, auto) + +(** more theorems about append **) + +lemma append_is_Nil_iff [simp]: + "xs:list(A) ==> (xs@ys = Nil) \ (xs=Nil & ys = Nil)" +by (erule list.induct, auto) + +lemma append_is_Nil_iff2 [simp]: + "xs:list(A) ==> (Nil = xs@ys) \ (xs=Nil & ys = Nil)" +by (erule list.induct, auto) + +lemma append_left_is_self_iff [simp]: + "xs:list(A) ==> (xs@ys = xs) \ (ys = Nil)" +by (erule list.induct, auto) + +lemma append_left_is_self_iff2 [simp]: + "xs:list(A) ==> (xs = xs@ys) \ (ys = Nil)" +by (erule list.induct, auto) + +(*TOO SLOW as a default simprule!*) +lemma append_left_is_Nil_iff [rule_format]: + "[| xs:list(A); ys:list(A); zs:list(A) |] ==> + length(ys)=length(zs) \ (xs@ys=zs \ (xs=Nil & ys=zs))" +apply (erule list.induct) +apply (auto simp add: length_app) +done + +(*TOO SLOW as a default simprule!*) +lemma append_left_is_Nil_iff2 [rule_format]: + "[| xs:list(A); ys:list(A); zs:list(A) |] ==> + length(ys)=length(zs) \ (zs=ys@xs \ (xs=Nil & ys=zs))" +apply (erule list.induct) +apply (auto simp add: length_app) +done + +lemma append_eq_append_iff [rule_format,simp]: + "xs:list(A) ==> \ys \ list(A). + length(xs)=length(ys) \ (xs@us = ys@vs) \ (xs=ys & us=vs)" +apply (erule list.induct) +apply (simp (no_asm_simp)) +apply clarify +apply (erule_tac a = ys in list.cases, auto) +done + +lemma append_eq_append [rule_format]: + "xs:list(A) ==> + \ys \ list(A). \us \ list(A). \vs \ list(A). + length(us) = length(vs) \ (xs@us = ys@vs) \ (xs=ys & us=vs)" +apply (induct_tac "xs") +apply (force simp add: length_app, clarify) +apply (erule_tac a = ys in list.cases, simp) +apply (subgoal_tac "Cons (a, l) @ us =vs") + apply (drule rev_iffD1 [OF _ append_left_is_Nil_iff], simp_all, blast) +done + +lemma append_eq_append_iff2 [simp]: + "[| xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs) |] + ==> xs@us = ys@vs \ (xs=ys & us=vs)" +apply (rule iffI) +apply (rule append_eq_append, auto) +done + +lemma append_self_iff [simp]: + "[| xs:list(A); ys:list(A); zs:list(A) |] ==> xs@ys=xs@zs \ ys=zs" +by simp + +lemma append_self_iff2 [simp]: + "[| xs:list(A); ys:list(A); zs:list(A) |] ==> ys@xs=zs@xs \ ys=zs" +by simp + +(* Can also be proved from append_eq_append_iff2, +but the proof requires two more hypotheses: x \ A and y \ A *) +lemma append1_eq_iff [rule_format,simp]: + "xs:list(A) ==> \ys \ list(A). xs@[x] = ys@[y] \ (xs = ys & x=y)" +apply (erule list.induct) + apply clarify + apply (erule list.cases) + apply simp_all +txt\Inductive step\ +apply clarify +apply (erule_tac a=ys in list.cases, simp_all) +done + + +lemma append_right_is_self_iff [simp]: + "[| xs:list(A); ys:list(A) |] ==> (xs@ys = ys) \ (xs=Nil)" +by (simp (no_asm_simp) add: append_left_is_Nil_iff) + +lemma append_right_is_self_iff2 [simp]: + "[| xs:list(A); ys:list(A) |] ==> (ys = xs@ys) \ (xs=Nil)" +apply (rule iffI) +apply (drule sym, auto) +done + +lemma hd_append [rule_format,simp]: + "xs:list(A) ==> xs \ Nil \ hd(xs @ ys) = hd(xs)" +by (induct_tac "xs", auto) + +lemma tl_append [rule_format,simp]: + "xs:list(A) ==> xs\Nil \ tl(xs @ ys) = tl(xs)@ys" +by (induct_tac "xs", auto) + +(** rev **) +lemma rev_is_Nil_iff [simp]: "xs:list(A) ==> (rev(xs) = Nil \ xs = Nil)" +by (erule list.induct, auto) + +lemma Nil_is_rev_iff [simp]: "xs:list(A) ==> (Nil = rev(xs) \ xs = Nil)" +by (erule list.induct, auto) + +lemma rev_is_rev_iff [rule_format,simp]: + "xs:list(A) ==> \ys \ list(A). rev(xs)=rev(ys) \ xs=ys" +apply (erule list.induct, force, clarify) +apply (erule_tac a = ys in list.cases, auto) +done + +lemma rev_list_elim [rule_format]: + "xs:list(A) ==> + (xs=Nil \ P) \ (\ys \ list(A). \y \ A. xs =ys@[y] \P)\P" +by (erule list_append_induct, auto) + + +(** more theorems about drop **) + +lemma length_drop [rule_format,simp]: + "n \ nat ==> \xs \ list(A). length(drop(n, xs)) = length(xs) #- n" +apply (erule nat_induct) +apply (auto elim: list.cases) +done + +lemma drop_all [rule_format,simp]: + "n \ nat ==> \xs \ list(A). length(xs) \ n \ drop(n, xs)=Nil" +apply (erule nat_induct) +apply (auto elim: list.cases) +done + +lemma drop_append [rule_format]: + "n \ nat ==> + \xs \ list(A). drop(n, xs@ys) = drop(n,xs) @ drop(n #- length(xs), ys)" +apply (induct_tac "n") +apply (auto elim: list.cases) +done + +lemma drop_drop: + "m \ nat ==> \xs \ list(A). \n \ nat. drop(n, drop(m, xs))=drop(n #+ m, xs)" +apply (induct_tac "m") +apply (auto elim: list.cases) +done + +(** take **) + +lemma take_0 [simp]: "xs:list(A) ==> take(0, xs) = Nil" +apply (unfold take_def) +apply (erule list.induct, auto) +done + +lemma take_succ_Cons [simp]: + "n \ nat ==> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))" +by (simp add: take_def) + +(* Needed for proving take_all *) +lemma take_Nil [simp]: "n \ nat ==> take(n, Nil) = Nil" +by (unfold take_def, auto) + +lemma take_all [rule_format,simp]: + "n \ nat ==> \xs \ list(A). length(xs) \ n \ take(n, xs) = xs" +apply (erule nat_induct) +apply (auto elim: list.cases) +done + +lemma take_type [rule_format,simp,TC]: + "xs:list(A) ==> \n \ nat. take(n, xs):list(A)" +apply (erule list.induct, simp, clarify) +apply (erule natE, auto) +done + +lemma take_append [rule_format,simp]: + "xs:list(A) ==> + \ys \ list(A). \n \ nat. take(n, xs @ ys) = + take(n, xs) @ take(n #- length(xs), ys)" +apply (erule list.induct, simp, clarify) +apply (erule natE, auto) +done + +lemma take_take [rule_format]: + "m \ nat ==> + \xs \ list(A). \n \ nat. take(n, take(m,xs))= take(min(n, m), xs)" +apply (induct_tac "m", auto) +apply (erule_tac a = xs in list.cases) +apply (auto simp add: take_Nil) +apply (erule_tac n=n in natE) +apply (auto intro: take_0 take_type) +done + +(** nth **) + +lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a" +by (simp add: nth_def) + +lemma nth_Cons [simp]: "n \ nat ==> nth(succ(n), Cons(a,l)) = nth(n,l)" +by (simp add: nth_def) + +lemma nth_empty [simp]: "nth(n, Nil) = 0" +by (simp add: nth_def) + +lemma nth_type [rule_format,simp,TC]: + "xs:list(A) ==> \n. n < length(xs) \ nth(n,xs) \ A" +apply (erule list.induct, simp, clarify) +apply (subgoal_tac "n \ nat") + apply (erule natE, auto dest!: le_in_nat) +done + +lemma nth_eq_0 [rule_format]: + "xs:list(A) ==> \n \ nat. length(xs) \ n \ nth(n,xs) = 0" +apply (erule list.induct, simp, clarify) +apply (erule natE, auto) +done + +lemma nth_append [rule_format]: + "xs:list(A) ==> + \n \ nat. nth(n, xs @ ys) = (if n < length(xs) then nth(n,xs) + else nth(n #- length(xs), ys))" +apply (induct_tac "xs", simp, clarify) +apply (erule natE, auto) +done + +lemma set_of_list_conv_nth: + "xs:list(A) + ==> set_of_list(xs) = {x \ A. \i\nat. i nat ==> + \xs \ list(A). (\ys \ list(A). k \ length(xs) \ k \ length(ys) \ + (\i \ nat. i nth(i,xs) = nth(i,ys))\ take(k,xs) = take(k,ys))" +apply (induct_tac "k") +apply (simp_all (no_asm_simp) add: lt_succ_eq_0_disj all_conj_distrib) +apply clarify +(*Both lists are non-empty*) +apply (erule_tac a=xs in list.cases, simp) +apply (erule_tac a=ys in list.cases, clarify) +apply (simp (no_asm_use) ) +apply clarify +apply (simp (no_asm_simp)) +apply (rule conjI, force) +apply (rename_tac y ys z zs) +apply (drule_tac x = zs and x1 = ys in bspec [THEN bspec], auto) +done + +lemma nth_equalityI [rule_format]: + "[| xs:list(A); ys:list(A); length(xs) = length(ys); + \i \ nat. i < length(xs) \ nth(i,xs) = nth(i,ys) |] + ==> xs = ys" +apply (subgoal_tac "length (xs) \ length (ys) ") +apply (cut_tac k="length(xs)" and xs=xs and ys=ys in nth_take_lemma) +apply (simp_all add: take_all) +done + +(*The famous take-lemma*) + +lemma take_equalityI [rule_format]: + "[| xs:list(A); ys:list(A); (\i \ nat. take(i, xs) = take(i,ys)) |] + ==> xs = ys" +apply (case_tac "length (xs) \ length (ys) ") +apply (drule_tac x = "length (ys) " in bspec) +apply (drule_tac [3] not_lt_imp_le) +apply (subgoal_tac [5] "length (ys) \ length (xs) ") +apply (rule_tac [6] j = "succ (length (ys))" in le_trans) +apply (rule_tac [6] leI) +apply (drule_tac [5] x = "length (xs) " in bspec) +apply (simp_all add: take_all) +done + +lemma nth_drop [rule_format]: + "n \ nat ==> \i \ nat. \xs \ list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)" +apply (induct_tac "n", simp_all, clarify) +apply (erule list.cases, auto) +done + +lemma take_succ [rule_format]: + "xs\list(A) + ==> \i. i < length(xs) \ take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]" +apply (induct_tac "xs", auto) +apply (subgoal_tac "i\nat") +apply (erule natE) +apply (auto simp add: le_in_nat) +done + +lemma take_add [rule_format]: + "[|xs\list(A); j\nat|] + ==> \i\nat. take(i #+ j, xs) = take(i,xs) @ take(j, drop(i,xs))" +apply (induct_tac "xs", simp_all, clarify) +apply (erule_tac n = i in natE, simp_all) +done + +lemma length_take: + "l\list(A) ==> \n\nat. length(take(n,l)) = min(n, length(l))" +apply (induct_tac "l", safe, simp_all) +apply (erule natE, simp_all) +done + +subsection\The function zip\ + +text\Crafty definition to eliminate a type argument\ + +consts + zip_aux :: "[i,i]=>i" + +primrec (*explicit lambda is required because both arguments of "un" vary*) + "zip_aux(B,[]) = + (\ys \ list(B). list_case([], %y l. [], ys))" + + "zip_aux(B,Cons(x,l)) = + (\ys \ list(B). + list_case(Nil, %y zs. Cons(, zip_aux(B,l)`zs), ys))" + +definition + zip :: "[i, i]=>i" where + "zip(xs, ys) == zip_aux(set_of_list(ys),xs)`ys" + + +(* zip equations *) + +lemma list_on_set_of_list: "xs \ list(A) ==> xs \ list(set_of_list(xs))" +apply (induct_tac xs, simp_all) +apply (blast intro: list_mono [THEN subsetD]) +done + +lemma zip_Nil [simp]: "ys:list(A) ==> zip(Nil, ys)=Nil" +apply (simp add: zip_def list_on_set_of_list [of _ A]) +apply (erule list.cases, simp_all) +done + +lemma zip_Nil2 [simp]: "xs:list(A) ==> zip(xs, Nil)=Nil" +apply (simp add: zip_def list_on_set_of_list [of _ A]) +apply (erule list.cases, simp_all) +done + +lemma zip_aux_unique [rule_format]: + "[|B<=C; xs \ list(A)|] + ==> \ys \ list(B). zip_aux(C,xs) ` ys = zip_aux(B,xs) ` ys" +apply (induct_tac xs) + apply simp_all + apply (blast intro: list_mono [THEN subsetD], clarify) +apply (erule_tac a=ys in list.cases, auto) +apply (blast intro: list_mono [THEN subsetD]) +done + +lemma zip_Cons_Cons [simp]: + "[| xs:list(A); ys:list(B); x \ A; y \ B |] ==> + zip(Cons(x,xs), Cons(y, ys)) = Cons(, zip(xs, ys))" +apply (simp add: zip_def, auto) +apply (rule zip_aux_unique, auto) +apply (simp add: list_on_set_of_list [of _ B]) +apply (blast intro: list_on_set_of_list list_mono [THEN subsetD]) +done + +lemma zip_type [rule_format,simp,TC]: + "xs:list(A) ==> \ys \ list(B). zip(xs, ys):list(A*B)" +apply (induct_tac "xs") +apply (simp (no_asm)) +apply clarify +apply (erule_tac a = ys in list.cases, auto) +done + +(* zip length *) +lemma length_zip [rule_format,simp]: + "xs:list(A) ==> \ys \ list(B). length(zip(xs,ys)) = + min(length(xs), length(ys))" +apply (unfold min_def) +apply (induct_tac "xs", simp_all, clarify) +apply (erule_tac a = ys in list.cases, auto) +done + +lemma zip_append1 [rule_format]: + "[| ys:list(A); zs:list(B) |] ==> + \xs \ list(A). zip(xs @ ys, zs) = + zip(xs, take(length(xs), zs)) @ zip(ys, drop(length(xs),zs))" +apply (induct_tac "zs", force, clarify) +apply (erule_tac a = xs in list.cases, simp_all) +done + +lemma zip_append2 [rule_format]: + "[| xs:list(A); zs:list(B) |] ==> \ys \ list(B). zip(xs, ys@zs) = + zip(take(length(ys), xs), ys) @ zip(drop(length(ys), xs), zs)" +apply (induct_tac "xs", force, clarify) +apply (erule_tac a = ys in list.cases, auto) +done + +lemma zip_append [simp]: + "[| length(xs) = length(us); length(ys) = length(vs); + xs:list(A); us:list(B); ys:list(A); vs:list(B) |] + ==> zip(xs@ys,us@vs) = zip(xs, us) @ zip(ys, vs)" +by (simp (no_asm_simp) add: zip_append1 drop_append diff_self_eq_0) + + +lemma zip_rev [rule_format,simp]: + "ys:list(B) ==> \xs \ list(A). + length(xs) = length(ys) \ zip(rev(xs), rev(ys)) = rev(zip(xs, ys))" +apply (induct_tac "ys", force, clarify) +apply (erule_tac a = xs in list.cases) +apply (auto simp add: length_rev) +done + +lemma nth_zip [rule_format,simp]: + "ys:list(B) ==> \i \ nat. \xs \ list(A). + i < length(xs) \ i < length(ys) \ + nth(i,zip(xs, ys)) = " +apply (induct_tac "ys", force, clarify) +apply (erule_tac a = xs in list.cases, simp) +apply (auto elim: natE) +done + +lemma set_of_list_zip [rule_format]: + "[| xs:list(A); ys:list(B); i \ nat |] + ==> set_of_list(zip(xs, ys)) = + {:A*B. \i\nat. i < min(length(xs), length(ys)) + & x = nth(i, xs) & y = nth(i, ys)}" +by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth) + +(** list_update **) + +lemma list_update_Nil [simp]: "i \ nat ==>list_update(Nil, i, v) = Nil" +by (unfold list_update_def, auto) + +lemma list_update_Cons_0 [simp]: "list_update(Cons(x, xs), 0, v)= Cons(v, xs)" +by (unfold list_update_def, auto) + +lemma list_update_Cons_succ [simp]: + "n \ nat ==> + list_update(Cons(x, xs), succ(n), v)= Cons(x, list_update(xs, n, v))" +apply (unfold list_update_def, auto) +done + +lemma list_update_type [rule_format,simp,TC]: + "[| xs:list(A); v \ A |] ==> \n \ nat. list_update(xs, n, v):list(A)" +apply (induct_tac "xs") +apply (simp (no_asm)) +apply clarify +apply (erule natE, auto) +done + +lemma length_list_update [rule_format,simp]: + "xs:list(A) ==> \i \ nat. length(list_update(xs, i, v))=length(xs)" +apply (induct_tac "xs") +apply (simp (no_asm)) +apply clarify +apply (erule natE, auto) +done + +lemma nth_list_update [rule_format]: + "[| xs:list(A) |] ==> \i \ nat. \j \ nat. i < length(xs) \ + nth(j, list_update(xs, i, x)) = (if i=j then x else nth(j, xs))" +apply (induct_tac "xs") + apply simp_all +apply clarify +apply (rename_tac i j) +apply (erule_tac n=i in natE) +apply (erule_tac [2] n=j in natE) +apply (erule_tac n=j in natE, simp_all, force) +done + +lemma nth_list_update_eq [simp]: + "[| i < length(xs); xs:list(A) |] ==> nth(i, list_update(xs, i,x)) = x" +by (simp (no_asm_simp) add: lt_length_in_nat nth_list_update) + + +lemma nth_list_update_neq [rule_format,simp]: + "xs:list(A) ==> + \i \ nat. \j \ nat. i \ j \ nth(j, list_update(xs,i,x)) = nth(j,xs)" +apply (induct_tac "xs") + apply (simp (no_asm)) +apply clarify +apply (erule natE) +apply (erule_tac [2] natE, simp_all) +apply (erule natE, simp_all) +done + +lemma list_update_overwrite [rule_format,simp]: + "xs:list(A) ==> \i \ nat. i < length(xs) + \ list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)" +apply (induct_tac "xs") + apply (simp (no_asm)) +apply clarify +apply (erule natE, auto) +done + +lemma list_update_same_conv [rule_format]: + "xs:list(A) ==> + \i \ nat. i < length(xs) \ + (list_update(xs, i, x) = xs) \ (nth(i, xs) = x)" +apply (induct_tac "xs") + apply (simp (no_asm)) +apply clarify +apply (erule natE, auto) +done + +lemma update_zip [rule_format]: + "ys:list(B) ==> + \i \ nat. \xy \ A*B. \xs \ list(A). + length(xs) = length(ys) \ + list_update(zip(xs, ys), i, xy) = zip(list_update(xs, i, fst(xy)), + list_update(ys, i, snd(xy)))" +apply (induct_tac "ys") + apply auto +apply (erule_tac a = xs in list.cases) +apply (auto elim: natE) +done + +lemma set_update_subset_cons [rule_format]: + "xs:list(A) ==> + \i \ nat. set_of_list(list_update(xs, i, x)) \ cons(x, set_of_list(xs))" +apply (induct_tac "xs") + apply simp +apply (rule ballI) +apply (erule natE, simp_all, auto) +done + +lemma set_of_list_update_subsetI: + "[| set_of_list(xs) \ A; xs:list(A); x \ A; i \ nat|] + ==> set_of_list(list_update(xs, i,x)) \ A" +apply (rule subset_trans) +apply (rule set_update_subset_cons, auto) +done + +(** upt **) + +lemma upt_rec: + "j \ nat ==> upt(i,j) = (if i i; j \ nat |] ==> upt(i,j) = Nil" +apply (subst upt_rec, auto) +apply (auto simp add: le_iff) +apply (drule lt_asym [THEN notE], auto) +done + +(*Only needed if upt_Suc is deleted from the simpset*) +lemma upt_succ_append: + "[| i \ j; j \ nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]" +by simp + +lemma upt_conv_Cons: + "[| i nat |] ==> upt(i,j) = Cons(i,upt(succ(i),j))" +apply (rule trans) +apply (rule upt_rec, auto) +done + +lemma upt_type [simp,TC]: "j \ nat ==> upt(i,j):list(nat)" +by (induct_tac "j", auto) + +(*LOOPS as a simprule, since j<=j*) +lemma upt_add_eq_append: + "[| i \ j; j \ nat; k \ nat |] ==> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)" +apply (induct_tac "k") +apply (auto simp add: app_assoc app_type) +apply (rule_tac j = j in le_trans, auto) +done + +lemma length_upt [simp]: "[| i \ nat; j \ nat |] ==>length(upt(i,j)) = j #- i" +apply (induct_tac "j") +apply (rule_tac [2] sym) +apply (auto dest!: not_lt_imp_le simp add: diff_succ diff_is_0_iff) +done + +lemma nth_upt [rule_format,simp]: + "[| i \ nat; j \ nat; k \ nat |] ==> i #+ k < j \ nth(k, upt(i,j)) = i #+ k" +apply (induct_tac "j", simp) +apply (simp add: nth_append le_iff) +apply (auto dest!: not_lt_imp_le + simp add: nth_append less_diff_conv add_commute) +done + +lemma take_upt [rule_format,simp]: + "[| m \ nat; n \ nat |] ==> + \i \ nat. i #+ m \ n \ take(m, upt(i,n)) = upt(i,i#+m)" +apply (induct_tac "m") +apply (simp (no_asm_simp) add: take_0) +apply clarify +apply (subst upt_rec, simp) +apply (rule sym) +apply (subst upt_rec, simp) +apply (simp_all del: upt.simps) +apply (rule_tac j = "succ (i #+ x) " in lt_trans2) +apply auto +done + +lemma map_succ_upt: + "[| m \ nat; n \ nat |] ==> map(succ, upt(m,n))= upt(succ(m), succ(n))" +apply (induct_tac "n") +apply (auto simp add: map_app_distrib) +done + +lemma nth_map [rule_format,simp]: + "xs:list(A) ==> + \n \ nat. n < length(xs) \ nth(n, map(f, xs)) = f(nth(n, xs))" +apply (induct_tac "xs", simp) +apply (rule ballI) +apply (induct_tac "n", auto) +done + +lemma nth_map_upt [rule_format]: + "[| m \ nat; n \ nat |] ==> + \i \ nat. i < n #- m \ nth(i, map(f, upt(m,n))) = f(m #+ i)" +apply (rule_tac n = m and m = n in diff_induct, typecheck, simp, simp) +apply (subst map_succ_upt [symmetric], simp_all, clarify) +apply (subgoal_tac "i < length (upt (0, x))") + prefer 2 + apply (simp add: less_diff_conv) + apply (rule_tac j = "succ (i #+ y) " in lt_trans2) + apply simp + apply simp +apply (subgoal_tac "i < length (upt (y, x))") + apply (simp_all add: add_commute less_diff_conv) +done + +(** sublist (a generalization of nth to sets) **) + +definition + sublist :: "[i, i] => i" where + "sublist(xs, A)== + map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))" + +lemma sublist_0 [simp]: "xs:list(A) ==>sublist(xs, 0) =Nil" +by (unfold sublist_def, auto) + +lemma sublist_Nil [simp]: "sublist(Nil, A) = Nil" +by (unfold sublist_def, auto) + +lemma sublist_shift_lemma: + "[| xs:list(B); i \ nat |] ==> + map(fst, filter(%p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) = + map(fst, filter(%p. snd(p):nat & snd(p) #+ i \ A, zip(xs,upt(0,length(xs)))))" +apply (erule list_append_induct) +apply (simp (no_asm_simp)) +apply (auto simp add: add_commute length_app filter_append map_app_distrib) +done + +lemma sublist_type [simp,TC]: + "xs:list(B) ==> sublist(xs, A):list(B)" +apply (unfold sublist_def) +apply (induct_tac "xs") +apply (auto simp add: filter_append map_app_distrib) +done + +lemma upt_add_eq_append2: + "[| i \ nat; j \ nat |] ==> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)" +by (simp add: upt_add_eq_append [of 0] nat_0_le) + +lemma sublist_append: + "[| xs:list(B); ys:list(B) |] ==> + sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j \ nat. j #+ length(xs): A})" +apply (unfold sublist_def) +apply (erule_tac l = ys in list_append_induct, simp) +apply (simp (no_asm_simp) add: upt_add_eq_append2 app_assoc [symmetric]) +apply (auto simp add: sublist_shift_lemma length_type map_app_distrib app_assoc) +apply (simp_all add: add_commute) +done + + +lemma sublist_Cons: + "[| xs:list(B); x \ B |] ==> + sublist(Cons(x, xs), A) = + (if 0 \ A then [x] else []) @ sublist(xs, {j \ nat. succ(j) \ A})" +apply (erule_tac l = xs in list_append_induct) +apply (simp (no_asm_simp) add: sublist_def) +apply (simp del: app_Cons add: app_Cons [symmetric] sublist_append, simp) +done + +lemma sublist_singleton [simp]: + "sublist([x], A) = (if 0 \ A then [x] else [])" +by (simp add: sublist_Cons) + +lemma sublist_upt_eq_take [rule_format, simp]: + "xs:list(A) ==> \n\nat. sublist(xs,n) = take(n,xs)" +apply (erule list.induct, simp) +apply (clarify ) +apply (erule natE) +apply (simp_all add: nat_eq_Collect_lt Ord_mem_iff_lt sublist_Cons) +done + +lemma sublist_Int_eq: + "xs \ list(B) ==> sublist(xs, A \ nat) = sublist(xs, A)" +apply (erule list.induct) +apply (simp_all add: sublist_Cons) +done + +text\Repetition of a List Element\ + +consts repeat :: "[i,i]=>i" +primrec + "repeat(a,0) = []" + + "repeat(a,succ(n)) = Cons(a,repeat(a,n))" + +lemma length_repeat: "n \ nat ==> length(repeat(a,n)) = n" +by (induct_tac n, auto) + +lemma repeat_succ_app: "n \ nat ==> repeat(a,succ(n)) = repeat(a,n) @ [a]" +apply (induct_tac n) +apply (simp_all del: app_Cons add: app_Cons [symmetric]) +done + +lemma repeat_type [TC]: "[|a \ A; n \ nat|] ==> repeat(a,n) \ list(A)" +by (induct_tac n, auto) + +end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/List_ZF.thy --- a/src/ZF/List_ZF.thy Sun Jun 24 11:41:41 2018 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1249 +0,0 @@ -(* Title: ZF/List_ZF.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1994 University of Cambridge -*) - -section\Lists in Zermelo-Fraenkel Set Theory\ - -theory List_ZF imports Datatype_ZF ArithSimp begin - -consts - list :: "i=>i" - -datatype - "list(A)" = Nil | Cons ("a \ A", "l \ list(A)") - - -syntax - "_Nil" :: i ("[]") - "_List" :: "is => i" ("[(_)]") - -translations - "[x, xs]" == "CONST Cons(x, [xs])" - "[x]" == "CONST Cons(x, [])" - "[]" == "CONST Nil" - - -consts - length :: "i=>i" - hd :: "i=>i" - tl :: "i=>i" - -primrec - "length([]) = 0" - "length(Cons(a,l)) = succ(length(l))" - -primrec - "hd([]) = 0" - "hd(Cons(a,l)) = a" - -primrec - "tl([]) = []" - "tl(Cons(a,l)) = l" - - -consts - map :: "[i=>i, i] => i" - set_of_list :: "i=>i" - app :: "[i,i]=>i" (infixr "@" 60) - -(*map is a binding operator -- it applies to meta-level functions, not -object-level functions. This simplifies the final form of term_rec_conv, -although complicating its derivation.*) -primrec - "map(f,[]) = []" - "map(f,Cons(a,l)) = Cons(f(a), map(f,l))" - -primrec - "set_of_list([]) = 0" - "set_of_list(Cons(a,l)) = cons(a, set_of_list(l))" - -primrec - app_Nil: "[] @ ys = ys" - app_Cons: "(Cons(a,l)) @ ys = Cons(a, l @ ys)" - - -consts - rev :: "i=>i" - flat :: "i=>i" - list_add :: "i=>i" - -primrec - "rev([]) = []" - "rev(Cons(a,l)) = rev(l) @ [a]" - -primrec - "flat([]) = []" - "flat(Cons(l,ls)) = l @ flat(ls)" - -primrec - "list_add([]) = 0" - "list_add(Cons(a,l)) = a #+ list_add(l)" - -consts - drop :: "[i,i]=>i" - -primrec - drop_0: "drop(0,l) = l" - drop_succ: "drop(succ(i), l) = tl (drop(i,l))" - - -(*** Thanks to Sidi Ehmety for the following ***) - -definition -(* Function `take' returns the first n elements of a list *) - take :: "[i,i]=>i" where - "take(n, as) == list_rec(\n\nat. [], - %a l r. \n\nat. nat_case([], %m. Cons(a, r`m), n), as)`n" - -definition - nth :: "[i, i]=>i" where - \ \returns the (n+1)th element of a list, or 0 if the - list is too short.\ - "nth(n, as) == list_rec(\n\nat. 0, - %a l r. \n\nat. nat_case(a, %m. r`m, n), as) ` n" - -definition - list_update :: "[i, i, i]=>i" where - "list_update(xs, i, v) == list_rec(\n\nat. Nil, - %u us vs. \n\nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i" - -consts - filter :: "[i=>o, i] => i" - upt :: "[i, i] =>i" - -primrec - "filter(P, Nil) = Nil" - "filter(P, Cons(x, xs)) = - (if P(x) then Cons(x, filter(P, xs)) else filter(P, xs))" - -primrec - "upt(i, 0) = Nil" - "upt(i, succ(j)) = (if i \ j then upt(i, j)@[j] else Nil)" - -definition - min :: "[i,i] =>i" where - "min(x, y) == (if x \ y then x else y)" - -definition - max :: "[i, i] =>i" where - "max(x, y) == (if x \ y then y else x)" - -(*** Aspects of the datatype definition ***) - -declare list.intros [simp,TC] - -(*An elimination rule, for type-checking*) -inductive_cases ConsE: "Cons(a,l) \ list(A)" - -lemma Cons_type_iff [simp]: "Cons(a,l) \ list(A) \ a \ A & l \ list(A)" -by (blast elim: ConsE) - -(*Proving freeness results*) -lemma Cons_iff: "Cons(a,l)=Cons(a',l') \ a=a' & l=l'" -by auto - -lemma Nil_Cons_iff: "~ Nil=Cons(a,l)" -by auto - -lemma list_unfold: "list(A) = {0} + (A * list(A))" -by (blast intro!: list.intros [unfolded list.con_defs] - elim: list.cases [unfolded list.con_defs]) - - -(** Lemmas to justify using "list" in other recursive type definitions **) - -lemma list_mono: "A<=B ==> list(A) \ list(B)" -apply (unfold list.defs ) -apply (rule lfp_mono) -apply (simp_all add: list.bnd_mono) -apply (assumption | rule univ_mono basic_monos)+ -done - -(*There is a similar proof by list induction.*) -lemma list_univ: "list(univ(A)) \ univ(A)" -apply (unfold list.defs list.con_defs) -apply (rule lfp_lowerbound) -apply (rule_tac [2] A_subset_univ [THEN univ_mono]) -apply (blast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ) -done - -(*These two theorems justify datatypes involving list(nat), list(A), ...*) -lemmas list_subset_univ = subset_trans [OF list_mono list_univ] - -lemma list_into_univ: "[| l \ list(A); A \ univ(B) |] ==> l \ univ(B)" -by (blast intro: list_subset_univ [THEN subsetD]) - -lemma list_case_type: - "[| l \ list(A); - c \ C(Nil); - !!x y. [| x \ A; y \ list(A) |] ==> h(x,y): C(Cons(x,y)) - |] ==> list_case(c,h,l) \ C(l)" -by (erule list.induct, auto) - -lemma list_0_triv: "list(0) = {Nil}" -apply (rule equalityI, auto) -apply (induct_tac x, auto) -done - - -(*** List functions ***) - -lemma tl_type: "l \ list(A) ==> tl(l) \ list(A)" -apply (induct_tac "l") -apply (simp_all (no_asm_simp) add: list.intros) -done - -(** drop **) - -lemma drop_Nil [simp]: "i \ nat ==> drop(i, Nil) = Nil" -apply (induct_tac "i") -apply (simp_all (no_asm_simp)) -done - -lemma drop_succ_Cons [simp]: "i \ nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)" -apply (rule sym) -apply (induct_tac "i") -apply (simp (no_asm)) -apply (simp (no_asm_simp)) -done - -lemma drop_type [simp,TC]: "[| i \ nat; l \ list(A) |] ==> drop(i,l) \ list(A)" -apply (induct_tac "i") -apply (simp_all (no_asm_simp) add: tl_type) -done - -declare drop_succ [simp del] - - -(** Type checking -- proved by induction, as usual **) - -lemma list_rec_type [TC]: - "[| l \ list(A); - c \ C(Nil); - !!x y r. [| x \ A; y \ list(A); r \ C(y) |] ==> h(x,y,r): C(Cons(x,y)) - |] ==> list_rec(c,h,l) \ C(l)" -by (induct_tac "l", auto) - -(** map **) - -lemma map_type [TC]: - "[| l \ list(A); !!x. x \ A ==> h(x): B |] ==> map(h,l) \ list(B)" -apply (simp add: map_list_def) -apply (typecheck add: list.intros list_rec_type, blast) -done - -lemma map_type2 [TC]: "l \ list(A) ==> map(h,l) \ list({h(u). u \ A})" -apply (erule map_type) -apply (erule RepFunI) -done - -(** length **) - -lemma length_type [TC]: "l \ list(A) ==> length(l) \ nat" -by (simp add: length_list_def) - -lemma lt_length_in_nat: - "[|x < length(xs); xs \ list(A)|] ==> x \ nat" -by (frule lt_nat_in_nat, typecheck) - -(** app **) - -lemma app_type [TC]: "[| xs: list(A); ys: list(A) |] ==> xs@ys \ list(A)" -by (simp add: app_list_def) - -(** rev **) - -lemma rev_type [TC]: "xs: list(A) ==> rev(xs) \ list(A)" -by (simp add: rev_list_def) - - -(** flat **) - -lemma flat_type [TC]: "ls: list(list(A)) ==> flat(ls) \ list(A)" -by (simp add: flat_list_def) - - -(** set_of_list **) - -lemma set_of_list_type [TC]: "l \ list(A) ==> set_of_list(l) \ Pow(A)" -apply (unfold set_of_list_list_def) -apply (erule list_rec_type, auto) -done - -lemma set_of_list_append: - "xs: list(A) ==> set_of_list (xs@ys) = set_of_list(xs) \ set_of_list(ys)" -apply (erule list.induct) -apply (simp_all (no_asm_simp) add: Un_cons) -done - - -(** list_add **) - -lemma list_add_type [TC]: "xs: list(nat) ==> list_add(xs) \ nat" -by (simp add: list_add_list_def) - - -(*** theorems about map ***) - -lemma map_ident [simp]: "l \ list(A) ==> map(%u. u, l) = l" -apply (induct_tac "l") -apply (simp_all (no_asm_simp)) -done - -lemma map_compose: "l \ list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)" -apply (induct_tac "l") -apply (simp_all (no_asm_simp)) -done - -lemma map_app_distrib: "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)" -apply (induct_tac "xs") -apply (simp_all (no_asm_simp)) -done - -lemma map_flat: "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))" -apply (induct_tac "ls") -apply (simp_all (no_asm_simp) add: map_app_distrib) -done - -lemma list_rec_map: - "l \ list(A) ==> - list_rec(c, d, map(h,l)) = - list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)" -apply (induct_tac "l") -apply (simp_all (no_asm_simp)) -done - -(** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **) - -(* @{term"c \ list(Collect(B,P)) ==> c \ list"} *) -lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD] - -lemma map_list_Collect: "l \ list({x \ A. h(x)=j(x)}) ==> map(h,l) = map(j,l)" -apply (induct_tac "l") -apply (simp_all (no_asm_simp)) -done - -(*** theorems about length ***) - -lemma length_map [simp]: "xs: list(A) ==> length(map(h,xs)) = length(xs)" -by (induct_tac "xs", simp_all) - -lemma length_app [simp]: - "[| xs: list(A); ys: list(A) |] - ==> length(xs@ys) = length(xs) #+ length(ys)" -by (induct_tac "xs", simp_all) - -lemma length_rev [simp]: "xs: list(A) ==> length(rev(xs)) = length(xs)" -apply (induct_tac "xs") -apply (simp_all (no_asm_simp) add: length_app) -done - -lemma length_flat: - "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))" -apply (induct_tac "ls") -apply (simp_all (no_asm_simp) add: length_app) -done - -(** Length and drop **) - -(*Lemma for the inductive step of drop_length*) -lemma drop_length_Cons [rule_format]: - "xs: list(A) ==> - \x. \z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)" -by (erule list.induct, simp_all) - -lemma drop_length [rule_format]: - "l \ list(A) ==> \i \ length(l). (\z zs. drop(i,l) = Cons(z,zs))" -apply (erule list.induct, simp_all, safe) -apply (erule drop_length_Cons) -apply (rule natE) -apply (erule Ord_trans [OF asm_rl length_type Ord_nat], assumption, simp_all) -apply (blast intro: succ_in_naturalD length_type) -done - - -(*** theorems about app ***) - -lemma app_right_Nil [simp]: "xs: list(A) ==> xs@Nil=xs" -by (erule list.induct, simp_all) - -lemma app_assoc: "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)" -by (induct_tac "xs", simp_all) - -lemma flat_app_distrib: "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)" -apply (induct_tac "ls") -apply (simp_all (no_asm_simp) add: app_assoc) -done - -(*** theorems about rev ***) - -lemma rev_map_distrib: "l \ list(A) ==> rev(map(h,l)) = map(h,rev(l))" -apply (induct_tac "l") -apply (simp_all (no_asm_simp) add: map_app_distrib) -done - -(*Simplifier needs the premises as assumptions because rewriting will not - instantiate the variable ?A in the rules' typing conditions; note that - rev_type does not instantiate ?A. Only the premises do. -*) -lemma rev_app_distrib: - "[| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)" -apply (erule list.induct) -apply (simp_all add: app_assoc) -done - -lemma rev_rev_ident [simp]: "l \ list(A) ==> rev(rev(l))=l" -apply (induct_tac "l") -apply (simp_all (no_asm_simp) add: rev_app_distrib) -done - -lemma rev_flat: "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))" -apply (induct_tac "ls") -apply (simp_all add: map_app_distrib flat_app_distrib rev_app_distrib) -done - - -(*** theorems about list_add ***) - -lemma list_add_app: - "[| xs: list(nat); ys: list(nat) |] - ==> list_add(xs@ys) = list_add(ys) #+ list_add(xs)" -apply (induct_tac "xs", simp_all) -done - -lemma list_add_rev: "l \ list(nat) ==> list_add(rev(l)) = list_add(l)" -apply (induct_tac "l") -apply (simp_all (no_asm_simp) add: list_add_app) -done - -lemma list_add_flat: - "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))" -apply (induct_tac "ls") -apply (simp_all (no_asm_simp) add: list_add_app) -done - -(** New induction rules **) - -lemma list_append_induct [case_names Nil snoc, consumes 1]: - "[| l \ list(A); - P(Nil); - !!x y. [| x \ A; y \ list(A); P(y) |] ==> P(y @ [x]) - |] ==> P(l)" -apply (subgoal_tac "P(rev(rev(l)))", simp) -apply (erule rev_type [THEN list.induct], simp_all) -done - -lemma list_complete_induct_lemma [rule_format]: - assumes ih: - "\l. [| l \ list(A); - \l' \ list(A). length(l') < length(l) \ P(l')|] - ==> P(l)" - shows "n \ nat ==> \l \ list(A). length(l) < n \ P(l)" -apply (induct_tac n, simp) -apply (blast intro: ih elim!: leE) -done - -theorem list_complete_induct: - "[| l \ list(A); - \l. [| l \ list(A); - \l' \ list(A). length(l') < length(l) \ P(l')|] - ==> P(l) - |] ==> P(l)" -apply (rule list_complete_induct_lemma [of A]) - prefer 4 apply (rule le_refl, simp) - apply blast - apply simp -apply assumption -done - - -(*** Thanks to Sidi Ehmety for these results about min, take, etc. ***) - -(** min FIXME: replace by Int! **) -(* Min theorems are also true for i, j ordinals *) -lemma min_sym: "[| i \ nat; j \ nat |] ==> min(i,j)=min(j,i)" -apply (unfold min_def) -apply (auto dest!: not_lt_imp_le dest: lt_not_sym intro: le_anti_sym) -done - -lemma min_type [simp,TC]: "[| i \ nat; j \ nat |] ==> min(i,j):nat" -by (unfold min_def, auto) - -lemma min_0 [simp]: "i \ nat ==> min(0,i) = 0" -apply (unfold min_def) -apply (auto dest: not_lt_imp_le) -done - -lemma min_02 [simp]: "i \ nat ==> min(i, 0) = 0" -apply (unfold min_def) -apply (auto dest: not_lt_imp_le) -done - -lemma lt_min_iff: "[| i \ nat; j \ nat; k \ nat |] ==> i i nat; j \ nat |] ==> min(succ(i), succ(j))= succ(min(i, j))" -apply (unfold min_def, auto) -done - -(*** more theorems about lists ***) - -(** filter **) - -lemma filter_append [simp]: - "xs:list(A) ==> filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)" -by (induct_tac "xs", auto) - -lemma filter_type [simp,TC]: "xs:list(A) ==> filter(P, xs):list(A)" -by (induct_tac "xs", auto) - -lemma length_filter: "xs:list(A) ==> length(filter(P, xs)) \ length(xs)" -apply (induct_tac "xs", auto) -apply (rule_tac j = "length (l) " in le_trans) -apply (auto simp add: le_iff) -done - -lemma filter_is_subset: "xs:list(A) ==> set_of_list(filter(P,xs)) \ set_of_list(xs)" -by (induct_tac "xs", auto) - -lemma filter_False [simp]: "xs:list(A) ==> filter(%p. False, xs) = Nil" -by (induct_tac "xs", auto) - -lemma filter_True [simp]: "xs:list(A) ==> filter(%p. True, xs) = xs" -by (induct_tac "xs", auto) - -(** length **) - -lemma length_is_0_iff [simp]: "xs:list(A) ==> length(xs)=0 \ xs=Nil" -by (erule list.induct, auto) - -lemma length_is_0_iff2 [simp]: "xs:list(A) ==> 0 = length(xs) \ xs=Nil" -by (erule list.induct, auto) - -lemma length_tl [simp]: "xs:list(A) ==> length(tl(xs)) = length(xs) #- 1" -by (erule list.induct, auto) - -lemma length_greater_0_iff: "xs:list(A) ==> 0 xs \ Nil" -by (erule list.induct, auto) - -lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) \ (\y ys. xs=Cons(y, ys) & length(ys)=n)" -by (erule list.induct, auto) - -(** more theorems about append **) - -lemma append_is_Nil_iff [simp]: - "xs:list(A) ==> (xs@ys = Nil) \ (xs=Nil & ys = Nil)" -by (erule list.induct, auto) - -lemma append_is_Nil_iff2 [simp]: - "xs:list(A) ==> (Nil = xs@ys) \ (xs=Nil & ys = Nil)" -by (erule list.induct, auto) - -lemma append_left_is_self_iff [simp]: - "xs:list(A) ==> (xs@ys = xs) \ (ys = Nil)" -by (erule list.induct, auto) - -lemma append_left_is_self_iff2 [simp]: - "xs:list(A) ==> (xs = xs@ys) \ (ys = Nil)" -by (erule list.induct, auto) - -(*TOO SLOW as a default simprule!*) -lemma append_left_is_Nil_iff [rule_format]: - "[| xs:list(A); ys:list(A); zs:list(A) |] ==> - length(ys)=length(zs) \ (xs@ys=zs \ (xs=Nil & ys=zs))" -apply (erule list.induct) -apply (auto simp add: length_app) -done - -(*TOO SLOW as a default simprule!*) -lemma append_left_is_Nil_iff2 [rule_format]: - "[| xs:list(A); ys:list(A); zs:list(A) |] ==> - length(ys)=length(zs) \ (zs=ys@xs \ (xs=Nil & ys=zs))" -apply (erule list.induct) -apply (auto simp add: length_app) -done - -lemma append_eq_append_iff [rule_format,simp]: - "xs:list(A) ==> \ys \ list(A). - length(xs)=length(ys) \ (xs@us = ys@vs) \ (xs=ys & us=vs)" -apply (erule list.induct) -apply (simp (no_asm_simp)) -apply clarify -apply (erule_tac a = ys in list.cases, auto) -done - -lemma append_eq_append [rule_format]: - "xs:list(A) ==> - \ys \ list(A). \us \ list(A). \vs \ list(A). - length(us) = length(vs) \ (xs@us = ys@vs) \ (xs=ys & us=vs)" -apply (induct_tac "xs") -apply (force simp add: length_app, clarify) -apply (erule_tac a = ys in list.cases, simp) -apply (subgoal_tac "Cons (a, l) @ us =vs") - apply (drule rev_iffD1 [OF _ append_left_is_Nil_iff], simp_all, blast) -done - -lemma append_eq_append_iff2 [simp]: - "[| xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs) |] - ==> xs@us = ys@vs \ (xs=ys & us=vs)" -apply (rule iffI) -apply (rule append_eq_append, auto) -done - -lemma append_self_iff [simp]: - "[| xs:list(A); ys:list(A); zs:list(A) |] ==> xs@ys=xs@zs \ ys=zs" -by simp - -lemma append_self_iff2 [simp]: - "[| xs:list(A); ys:list(A); zs:list(A) |] ==> ys@xs=zs@xs \ ys=zs" -by simp - -(* Can also be proved from append_eq_append_iff2, -but the proof requires two more hypotheses: x \ A and y \ A *) -lemma append1_eq_iff [rule_format,simp]: - "xs:list(A) ==> \ys \ list(A). xs@[x] = ys@[y] \ (xs = ys & x=y)" -apply (erule list.induct) - apply clarify - apply (erule list.cases) - apply simp_all -txt\Inductive step\ -apply clarify -apply (erule_tac a=ys in list.cases, simp_all) -done - - -lemma append_right_is_self_iff [simp]: - "[| xs:list(A); ys:list(A) |] ==> (xs@ys = ys) \ (xs=Nil)" -by (simp (no_asm_simp) add: append_left_is_Nil_iff) - -lemma append_right_is_self_iff2 [simp]: - "[| xs:list(A); ys:list(A) |] ==> (ys = xs@ys) \ (xs=Nil)" -apply (rule iffI) -apply (drule sym, auto) -done - -lemma hd_append [rule_format,simp]: - "xs:list(A) ==> xs \ Nil \ hd(xs @ ys) = hd(xs)" -by (induct_tac "xs", auto) - -lemma tl_append [rule_format,simp]: - "xs:list(A) ==> xs\Nil \ tl(xs @ ys) = tl(xs)@ys" -by (induct_tac "xs", auto) - -(** rev **) -lemma rev_is_Nil_iff [simp]: "xs:list(A) ==> (rev(xs) = Nil \ xs = Nil)" -by (erule list.induct, auto) - -lemma Nil_is_rev_iff [simp]: "xs:list(A) ==> (Nil = rev(xs) \ xs = Nil)" -by (erule list.induct, auto) - -lemma rev_is_rev_iff [rule_format,simp]: - "xs:list(A) ==> \ys \ list(A). rev(xs)=rev(ys) \ xs=ys" -apply (erule list.induct, force, clarify) -apply (erule_tac a = ys in list.cases, auto) -done - -lemma rev_list_elim [rule_format]: - "xs:list(A) ==> - (xs=Nil \ P) \ (\ys \ list(A). \y \ A. xs =ys@[y] \P)\P" -by (erule list_append_induct, auto) - - -(** more theorems about drop **) - -lemma length_drop [rule_format,simp]: - "n \ nat ==> \xs \ list(A). length(drop(n, xs)) = length(xs) #- n" -apply (erule nat_induct) -apply (auto elim: list.cases) -done - -lemma drop_all [rule_format,simp]: - "n \ nat ==> \xs \ list(A). length(xs) \ n \ drop(n, xs)=Nil" -apply (erule nat_induct) -apply (auto elim: list.cases) -done - -lemma drop_append [rule_format]: - "n \ nat ==> - \xs \ list(A). drop(n, xs@ys) = drop(n,xs) @ drop(n #- length(xs), ys)" -apply (induct_tac "n") -apply (auto elim: list.cases) -done - -lemma drop_drop: - "m \ nat ==> \xs \ list(A). \n \ nat. drop(n, drop(m, xs))=drop(n #+ m, xs)" -apply (induct_tac "m") -apply (auto elim: list.cases) -done - -(** take **) - -lemma take_0 [simp]: "xs:list(A) ==> take(0, xs) = Nil" -apply (unfold take_def) -apply (erule list.induct, auto) -done - -lemma take_succ_Cons [simp]: - "n \ nat ==> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))" -by (simp add: take_def) - -(* Needed for proving take_all *) -lemma take_Nil [simp]: "n \ nat ==> take(n, Nil) = Nil" -by (unfold take_def, auto) - -lemma take_all [rule_format,simp]: - "n \ nat ==> \xs \ list(A). length(xs) \ n \ take(n, xs) = xs" -apply (erule nat_induct) -apply (auto elim: list.cases) -done - -lemma take_type [rule_format,simp,TC]: - "xs:list(A) ==> \n \ nat. take(n, xs):list(A)" -apply (erule list.induct, simp, clarify) -apply (erule natE, auto) -done - -lemma take_append [rule_format,simp]: - "xs:list(A) ==> - \ys \ list(A). \n \ nat. take(n, xs @ ys) = - take(n, xs) @ take(n #- length(xs), ys)" -apply (erule list.induct, simp, clarify) -apply (erule natE, auto) -done - -lemma take_take [rule_format]: - "m \ nat ==> - \xs \ list(A). \n \ nat. take(n, take(m,xs))= take(min(n, m), xs)" -apply (induct_tac "m", auto) -apply (erule_tac a = xs in list.cases) -apply (auto simp add: take_Nil) -apply (erule_tac n=n in natE) -apply (auto intro: take_0 take_type) -done - -(** nth **) - -lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a" -by (simp add: nth_def) - -lemma nth_Cons [simp]: "n \ nat ==> nth(succ(n), Cons(a,l)) = nth(n,l)" -by (simp add: nth_def) - -lemma nth_empty [simp]: "nth(n, Nil) = 0" -by (simp add: nth_def) - -lemma nth_type [rule_format,simp,TC]: - "xs:list(A) ==> \n. n < length(xs) \ nth(n,xs) \ A" -apply (erule list.induct, simp, clarify) -apply (subgoal_tac "n \ nat") - apply (erule natE, auto dest!: le_in_nat) -done - -lemma nth_eq_0 [rule_format]: - "xs:list(A) ==> \n \ nat. length(xs) \ n \ nth(n,xs) = 0" -apply (erule list.induct, simp, clarify) -apply (erule natE, auto) -done - -lemma nth_append [rule_format]: - "xs:list(A) ==> - \n \ nat. nth(n, xs @ ys) = (if n < length(xs) then nth(n,xs) - else nth(n #- length(xs), ys))" -apply (induct_tac "xs", simp, clarify) -apply (erule natE, auto) -done - -lemma set_of_list_conv_nth: - "xs:list(A) - ==> set_of_list(xs) = {x \ A. \i\nat. i nat ==> - \xs \ list(A). (\ys \ list(A). k \ length(xs) \ k \ length(ys) \ - (\i \ nat. i nth(i,xs) = nth(i,ys))\ take(k,xs) = take(k,ys))" -apply (induct_tac "k") -apply (simp_all (no_asm_simp) add: lt_succ_eq_0_disj all_conj_distrib) -apply clarify -(*Both lists are non-empty*) -apply (erule_tac a=xs in list.cases, simp) -apply (erule_tac a=ys in list.cases, clarify) -apply (simp (no_asm_use) ) -apply clarify -apply (simp (no_asm_simp)) -apply (rule conjI, force) -apply (rename_tac y ys z zs) -apply (drule_tac x = zs and x1 = ys in bspec [THEN bspec], auto) -done - -lemma nth_equalityI [rule_format]: - "[| xs:list(A); ys:list(A); length(xs) = length(ys); - \i \ nat. i < length(xs) \ nth(i,xs) = nth(i,ys) |] - ==> xs = ys" -apply (subgoal_tac "length (xs) \ length (ys) ") -apply (cut_tac k="length(xs)" and xs=xs and ys=ys in nth_take_lemma) -apply (simp_all add: take_all) -done - -(*The famous take-lemma*) - -lemma take_equalityI [rule_format]: - "[| xs:list(A); ys:list(A); (\i \ nat. take(i, xs) = take(i,ys)) |] - ==> xs = ys" -apply (case_tac "length (xs) \ length (ys) ") -apply (drule_tac x = "length (ys) " in bspec) -apply (drule_tac [3] not_lt_imp_le) -apply (subgoal_tac [5] "length (ys) \ length (xs) ") -apply (rule_tac [6] j = "succ (length (ys))" in le_trans) -apply (rule_tac [6] leI) -apply (drule_tac [5] x = "length (xs) " in bspec) -apply (simp_all add: take_all) -done - -lemma nth_drop [rule_format]: - "n \ nat ==> \i \ nat. \xs \ list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)" -apply (induct_tac "n", simp_all, clarify) -apply (erule list.cases, auto) -done - -lemma take_succ [rule_format]: - "xs\list(A) - ==> \i. i < length(xs) \ take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]" -apply (induct_tac "xs", auto) -apply (subgoal_tac "i\nat") -apply (erule natE) -apply (auto simp add: le_in_nat) -done - -lemma take_add [rule_format]: - "[|xs\list(A); j\nat|] - ==> \i\nat. take(i #+ j, xs) = take(i,xs) @ take(j, drop(i,xs))" -apply (induct_tac "xs", simp_all, clarify) -apply (erule_tac n = i in natE, simp_all) -done - -lemma length_take: - "l\list(A) ==> \n\nat. length(take(n,l)) = min(n, length(l))" -apply (induct_tac "l", safe, simp_all) -apply (erule natE, simp_all) -done - -subsection\The function zip\ - -text\Crafty definition to eliminate a type argument\ - -consts - zip_aux :: "[i,i]=>i" - -primrec (*explicit lambda is required because both arguments of "un" vary*) - "zip_aux(B,[]) = - (\ys \ list(B). list_case([], %y l. [], ys))" - - "zip_aux(B,Cons(x,l)) = - (\ys \ list(B). - list_case(Nil, %y zs. Cons(, zip_aux(B,l)`zs), ys))" - -definition - zip :: "[i, i]=>i" where - "zip(xs, ys) == zip_aux(set_of_list(ys),xs)`ys" - - -(* zip equations *) - -lemma list_on_set_of_list: "xs \ list(A) ==> xs \ list(set_of_list(xs))" -apply (induct_tac xs, simp_all) -apply (blast intro: list_mono [THEN subsetD]) -done - -lemma zip_Nil [simp]: "ys:list(A) ==> zip(Nil, ys)=Nil" -apply (simp add: zip_def list_on_set_of_list [of _ A]) -apply (erule list.cases, simp_all) -done - -lemma zip_Nil2 [simp]: "xs:list(A) ==> zip(xs, Nil)=Nil" -apply (simp add: zip_def list_on_set_of_list [of _ A]) -apply (erule list.cases, simp_all) -done - -lemma zip_aux_unique [rule_format]: - "[|B<=C; xs \ list(A)|] - ==> \ys \ list(B). zip_aux(C,xs) ` ys = zip_aux(B,xs) ` ys" -apply (induct_tac xs) - apply simp_all - apply (blast intro: list_mono [THEN subsetD], clarify) -apply (erule_tac a=ys in list.cases, auto) -apply (blast intro: list_mono [THEN subsetD]) -done - -lemma zip_Cons_Cons [simp]: - "[| xs:list(A); ys:list(B); x \ A; y \ B |] ==> - zip(Cons(x,xs), Cons(y, ys)) = Cons(, zip(xs, ys))" -apply (simp add: zip_def, auto) -apply (rule zip_aux_unique, auto) -apply (simp add: list_on_set_of_list [of _ B]) -apply (blast intro: list_on_set_of_list list_mono [THEN subsetD]) -done - -lemma zip_type [rule_format,simp,TC]: - "xs:list(A) ==> \ys \ list(B). zip(xs, ys):list(A*B)" -apply (induct_tac "xs") -apply (simp (no_asm)) -apply clarify -apply (erule_tac a = ys in list.cases, auto) -done - -(* zip length *) -lemma length_zip [rule_format,simp]: - "xs:list(A) ==> \ys \ list(B). length(zip(xs,ys)) = - min(length(xs), length(ys))" -apply (unfold min_def) -apply (induct_tac "xs", simp_all, clarify) -apply (erule_tac a = ys in list.cases, auto) -done - -lemma zip_append1 [rule_format]: - "[| ys:list(A); zs:list(B) |] ==> - \xs \ list(A). zip(xs @ ys, zs) = - zip(xs, take(length(xs), zs)) @ zip(ys, drop(length(xs),zs))" -apply (induct_tac "zs", force, clarify) -apply (erule_tac a = xs in list.cases, simp_all) -done - -lemma zip_append2 [rule_format]: - "[| xs:list(A); zs:list(B) |] ==> \ys \ list(B). zip(xs, ys@zs) = - zip(take(length(ys), xs), ys) @ zip(drop(length(ys), xs), zs)" -apply (induct_tac "xs", force, clarify) -apply (erule_tac a = ys in list.cases, auto) -done - -lemma zip_append [simp]: - "[| length(xs) = length(us); length(ys) = length(vs); - xs:list(A); us:list(B); ys:list(A); vs:list(B) |] - ==> zip(xs@ys,us@vs) = zip(xs, us) @ zip(ys, vs)" -by (simp (no_asm_simp) add: zip_append1 drop_append diff_self_eq_0) - - -lemma zip_rev [rule_format,simp]: - "ys:list(B) ==> \xs \ list(A). - length(xs) = length(ys) \ zip(rev(xs), rev(ys)) = rev(zip(xs, ys))" -apply (induct_tac "ys", force, clarify) -apply (erule_tac a = xs in list.cases) -apply (auto simp add: length_rev) -done - -lemma nth_zip [rule_format,simp]: - "ys:list(B) ==> \i \ nat. \xs \ list(A). - i < length(xs) \ i < length(ys) \ - nth(i,zip(xs, ys)) = " -apply (induct_tac "ys", force, clarify) -apply (erule_tac a = xs in list.cases, simp) -apply (auto elim: natE) -done - -lemma set_of_list_zip [rule_format]: - "[| xs:list(A); ys:list(B); i \ nat |] - ==> set_of_list(zip(xs, ys)) = - {:A*B. \i\nat. i < min(length(xs), length(ys)) - & x = nth(i, xs) & y = nth(i, ys)}" -by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth) - -(** list_update **) - -lemma list_update_Nil [simp]: "i \ nat ==>list_update(Nil, i, v) = Nil" -by (unfold list_update_def, auto) - -lemma list_update_Cons_0 [simp]: "list_update(Cons(x, xs), 0, v)= Cons(v, xs)" -by (unfold list_update_def, auto) - -lemma list_update_Cons_succ [simp]: - "n \ nat ==> - list_update(Cons(x, xs), succ(n), v)= Cons(x, list_update(xs, n, v))" -apply (unfold list_update_def, auto) -done - -lemma list_update_type [rule_format,simp,TC]: - "[| xs:list(A); v \ A |] ==> \n \ nat. list_update(xs, n, v):list(A)" -apply (induct_tac "xs") -apply (simp (no_asm)) -apply clarify -apply (erule natE, auto) -done - -lemma length_list_update [rule_format,simp]: - "xs:list(A) ==> \i \ nat. length(list_update(xs, i, v))=length(xs)" -apply (induct_tac "xs") -apply (simp (no_asm)) -apply clarify -apply (erule natE, auto) -done - -lemma nth_list_update [rule_format]: - "[| xs:list(A) |] ==> \i \ nat. \j \ nat. i < length(xs) \ - nth(j, list_update(xs, i, x)) = (if i=j then x else nth(j, xs))" -apply (induct_tac "xs") - apply simp_all -apply clarify -apply (rename_tac i j) -apply (erule_tac n=i in natE) -apply (erule_tac [2] n=j in natE) -apply (erule_tac n=j in natE, simp_all, force) -done - -lemma nth_list_update_eq [simp]: - "[| i < length(xs); xs:list(A) |] ==> nth(i, list_update(xs, i,x)) = x" -by (simp (no_asm_simp) add: lt_length_in_nat nth_list_update) - - -lemma nth_list_update_neq [rule_format,simp]: - "xs:list(A) ==> - \i \ nat. \j \ nat. i \ j \ nth(j, list_update(xs,i,x)) = nth(j,xs)" -apply (induct_tac "xs") - apply (simp (no_asm)) -apply clarify -apply (erule natE) -apply (erule_tac [2] natE, simp_all) -apply (erule natE, simp_all) -done - -lemma list_update_overwrite [rule_format,simp]: - "xs:list(A) ==> \i \ nat. i < length(xs) - \ list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)" -apply (induct_tac "xs") - apply (simp (no_asm)) -apply clarify -apply (erule natE, auto) -done - -lemma list_update_same_conv [rule_format]: - "xs:list(A) ==> - \i \ nat. i < length(xs) \ - (list_update(xs, i, x) = xs) \ (nth(i, xs) = x)" -apply (induct_tac "xs") - apply (simp (no_asm)) -apply clarify -apply (erule natE, auto) -done - -lemma update_zip [rule_format]: - "ys:list(B) ==> - \i \ nat. \xy \ A*B. \xs \ list(A). - length(xs) = length(ys) \ - list_update(zip(xs, ys), i, xy) = zip(list_update(xs, i, fst(xy)), - list_update(ys, i, snd(xy)))" -apply (induct_tac "ys") - apply auto -apply (erule_tac a = xs in list.cases) -apply (auto elim: natE) -done - -lemma set_update_subset_cons [rule_format]: - "xs:list(A) ==> - \i \ nat. set_of_list(list_update(xs, i, x)) \ cons(x, set_of_list(xs))" -apply (induct_tac "xs") - apply simp -apply (rule ballI) -apply (erule natE, simp_all, auto) -done - -lemma set_of_list_update_subsetI: - "[| set_of_list(xs) \ A; xs:list(A); x \ A; i \ nat|] - ==> set_of_list(list_update(xs, i,x)) \ A" -apply (rule subset_trans) -apply (rule set_update_subset_cons, auto) -done - -(** upt **) - -lemma upt_rec: - "j \ nat ==> upt(i,j) = (if i i; j \ nat |] ==> upt(i,j) = Nil" -apply (subst upt_rec, auto) -apply (auto simp add: le_iff) -apply (drule lt_asym [THEN notE], auto) -done - -(*Only needed if upt_Suc is deleted from the simpset*) -lemma upt_succ_append: - "[| i \ j; j \ nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]" -by simp - -lemma upt_conv_Cons: - "[| i nat |] ==> upt(i,j) = Cons(i,upt(succ(i),j))" -apply (rule trans) -apply (rule upt_rec, auto) -done - -lemma upt_type [simp,TC]: "j \ nat ==> upt(i,j):list(nat)" -by (induct_tac "j", auto) - -(*LOOPS as a simprule, since j<=j*) -lemma upt_add_eq_append: - "[| i \ j; j \ nat; k \ nat |] ==> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)" -apply (induct_tac "k") -apply (auto simp add: app_assoc app_type) -apply (rule_tac j = j in le_trans, auto) -done - -lemma length_upt [simp]: "[| i \ nat; j \ nat |] ==>length(upt(i,j)) = j #- i" -apply (induct_tac "j") -apply (rule_tac [2] sym) -apply (auto dest!: not_lt_imp_le simp add: diff_succ diff_is_0_iff) -done - -lemma nth_upt [rule_format,simp]: - "[| i \ nat; j \ nat; k \ nat |] ==> i #+ k < j \ nth(k, upt(i,j)) = i #+ k" -apply (induct_tac "j", simp) -apply (simp add: nth_append le_iff) -apply (auto dest!: not_lt_imp_le - simp add: nth_append less_diff_conv add_commute) -done - -lemma take_upt [rule_format,simp]: - "[| m \ nat; n \ nat |] ==> - \i \ nat. i #+ m \ n \ take(m, upt(i,n)) = upt(i,i#+m)" -apply (induct_tac "m") -apply (simp (no_asm_simp) add: take_0) -apply clarify -apply (subst upt_rec, simp) -apply (rule sym) -apply (subst upt_rec, simp) -apply (simp_all del: upt.simps) -apply (rule_tac j = "succ (i #+ x) " in lt_trans2) -apply auto -done - -lemma map_succ_upt: - "[| m \ nat; n \ nat |] ==> map(succ, upt(m,n))= upt(succ(m), succ(n))" -apply (induct_tac "n") -apply (auto simp add: map_app_distrib) -done - -lemma nth_map [rule_format,simp]: - "xs:list(A) ==> - \n \ nat. n < length(xs) \ nth(n, map(f, xs)) = f(nth(n, xs))" -apply (induct_tac "xs", simp) -apply (rule ballI) -apply (induct_tac "n", auto) -done - -lemma nth_map_upt [rule_format]: - "[| m \ nat; n \ nat |] ==> - \i \ nat. i < n #- m \ nth(i, map(f, upt(m,n))) = f(m #+ i)" -apply (rule_tac n = m and m = n in diff_induct, typecheck, simp, simp) -apply (subst map_succ_upt [symmetric], simp_all, clarify) -apply (subgoal_tac "i < length (upt (0, x))") - prefer 2 - apply (simp add: less_diff_conv) - apply (rule_tac j = "succ (i #+ y) " in lt_trans2) - apply simp - apply simp -apply (subgoal_tac "i < length (upt (y, x))") - apply (simp_all add: add_commute less_diff_conv) -done - -(** sublist (a generalization of nth to sets) **) - -definition - sublist :: "[i, i] => i" where - "sublist(xs, A)== - map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))" - -lemma sublist_0 [simp]: "xs:list(A) ==>sublist(xs, 0) =Nil" -by (unfold sublist_def, auto) - -lemma sublist_Nil [simp]: "sublist(Nil, A) = Nil" -by (unfold sublist_def, auto) - -lemma sublist_shift_lemma: - "[| xs:list(B); i \ nat |] ==> - map(fst, filter(%p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) = - map(fst, filter(%p. snd(p):nat & snd(p) #+ i \ A, zip(xs,upt(0,length(xs)))))" -apply (erule list_append_induct) -apply (simp (no_asm_simp)) -apply (auto simp add: add_commute length_app filter_append map_app_distrib) -done - -lemma sublist_type [simp,TC]: - "xs:list(B) ==> sublist(xs, A):list(B)" -apply (unfold sublist_def) -apply (induct_tac "xs") -apply (auto simp add: filter_append map_app_distrib) -done - -lemma upt_add_eq_append2: - "[| i \ nat; j \ nat |] ==> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)" -by (simp add: upt_add_eq_append [of 0] nat_0_le) - -lemma sublist_append: - "[| xs:list(B); ys:list(B) |] ==> - sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j \ nat. j #+ length(xs): A})" -apply (unfold sublist_def) -apply (erule_tac l = ys in list_append_induct, simp) -apply (simp (no_asm_simp) add: upt_add_eq_append2 app_assoc [symmetric]) -apply (auto simp add: sublist_shift_lemma length_type map_app_distrib app_assoc) -apply (simp_all add: add_commute) -done - - -lemma sublist_Cons: - "[| xs:list(B); x \ B |] ==> - sublist(Cons(x, xs), A) = - (if 0 \ A then [x] else []) @ sublist(xs, {j \ nat. succ(j) \ A})" -apply (erule_tac l = xs in list_append_induct) -apply (simp (no_asm_simp) add: sublist_def) -apply (simp del: app_Cons add: app_Cons [symmetric] sublist_append, simp) -done - -lemma sublist_singleton [simp]: - "sublist([x], A) = (if 0 \ A then [x] else [])" -by (simp add: sublist_Cons) - -lemma sublist_upt_eq_take [rule_format, simp]: - "xs:list(A) ==> \n\nat. sublist(xs,n) = take(n,xs)" -apply (erule list.induct, simp) -apply (clarify ) -apply (erule natE) -apply (simp_all add: nat_eq_Collect_lt Ord_mem_iff_lt sublist_Cons) -done - -lemma sublist_Int_eq: - "xs \ list(B) ==> sublist(xs, A \ nat) = sublist(xs, A)" -apply (erule list.induct) -apply (simp_all add: sublist_Cons) -done - -text\Repetition of a List Element\ - -consts repeat :: "[i,i]=>i" -primrec - "repeat(a,0) = []" - - "repeat(a,succ(n)) = Cons(a,repeat(a,n))" - -lemma length_repeat: "n \ nat ==> length(repeat(a,n)) = n" -by (induct_tac n, auto) - -lemma repeat_succ_app: "n \ nat ==> repeat(a,succ(n)) = repeat(a,n) @ [a]" -apply (induct_tac n) -apply (simp_all del: app_Cons add: app_Cons [symmetric]) -done - -lemma repeat_type [TC]: "[|a \ A; n \ nat|] ==> repeat(a,n) \ list(A)" -by (induct_tac n, auto) - -end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Nat.thy Sun Jun 24 15:57:48 2018 +0200 @@ -0,0 +1,302 @@ +(* Title: ZF/Nat.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge +*) + +section\The Natural numbers As a Least Fixed Point\ + +theory Nat imports OrdQuant Bool begin + +definition + nat :: i where + "nat == lfp(Inf, %X. {0} \ {succ(i). i \ X})" + +definition + quasinat :: "i => o" where + "quasinat(n) == n=0 | (\m. n = succ(m))" + +definition + (*Has an unconditional succ case, which is used in "recursor" below.*) + nat_case :: "[i, i=>i, i]=>i" where + "nat_case(a,b,k) == THE y. k=0 & y=a | (\x. k=succ(x) & y=b(x))" + +definition + nat_rec :: "[i, i, [i,i]=>i]=>i" where + "nat_rec(k,a,b) == + wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))" + + (*Internalized relations on the naturals*) + +definition + Le :: i where + "Le == {:nat*nat. x \ y}" + +definition + Lt :: i where + "Lt == {:nat*nat. x < y}" + +definition + Ge :: i where + "Ge == {:nat*nat. y \ x}" + +definition + Gt :: i where + "Gt == {:nat*nat. y < x}" + +definition + greater_than :: "i=>i" where + "greater_than(n) == {i \ nat. n < i}" + +text\No need for a less-than operator: a natural number is its list of +predecessors!\ + + +lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} \ {succ(i). i \ X})" +apply (rule bnd_monoI) +apply (cut_tac infinity, blast, blast) +done + +(* @{term"nat = {0} \ {succ(x). x \ nat}"} *) +lemmas nat_unfold = nat_bnd_mono [THEN nat_def [THEN def_lfp_unfold]] + +(** Type checking of 0 and successor **) + +lemma nat_0I [iff,TC]: "0 \ nat" +apply (subst nat_unfold) +apply (rule singletonI [THEN UnI1]) +done + +lemma nat_succI [intro!,TC]: "n \ nat ==> succ(n) \ nat" +apply (subst nat_unfold) +apply (erule RepFunI [THEN UnI2]) +done + +lemma nat_1I [iff,TC]: "1 \ nat" +by (rule nat_0I [THEN nat_succI]) + +lemma nat_2I [iff,TC]: "2 \ nat" +by (rule nat_1I [THEN nat_succI]) + +lemma bool_subset_nat: "bool \ nat" +by (blast elim!: boolE) + +lemmas bool_into_nat = bool_subset_nat [THEN subsetD] + + +subsection\Injectivity Properties and Induction\ + +(*Mathematical induction*) +lemma nat_induct [case_names 0 succ, induct set: nat]: + "[| n \ nat; P(0); !!x. [| x \ nat; P(x) |] ==> P(succ(x)) |] ==> P(n)" +by (erule def_induct [OF nat_def nat_bnd_mono], blast) + +lemma natE: + assumes "n \ nat" + obtains ("0") "n=0" | (succ) x where "x \ nat" "n=succ(x)" +using assms +by (rule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE]) auto + +lemma nat_into_Ord [simp]: "n \ nat ==> Ord(n)" +by (erule nat_induct, auto) + +(* @{term"i \ nat ==> 0 \ i"}; same thing as @{term"0 nat ==> i \ i"}; same thing as @{term"i nat ==> ~ Limit(a)" +by (induct a rule: nat_induct, auto) + +lemma succ_natD: "succ(i): nat ==> i \ nat" +by (rule Ord_trans [OF succI1], auto) + +lemma nat_succ_iff [iff]: "succ(n): nat \ n \ nat" +by (blast dest!: succ_natD) + +lemma nat_le_Limit: "Limit(i) ==> nat \ i" +apply (rule subset_imp_le) +apply (simp_all add: Limit_is_Ord) +apply (rule subsetI) +apply (erule nat_induct) + apply (erule Limit_has_0 [THEN ltD]) +apply (blast intro: Limit_has_succ [THEN ltD] ltI Limit_is_Ord) +done + +(* [| succ(i): k; k \ nat |] ==> i \ k *) +lemmas succ_in_naturalD = Ord_trans [OF succI1 _ nat_into_Ord] + +lemma lt_nat_in_nat: "[| m nat |] ==> m \ nat" +apply (erule ltE) +apply (erule Ord_trans, assumption, simp) +done + +lemma le_in_nat: "[| m \ n; n \ nat |] ==> m \ nat" +by (blast dest!: lt_nat_in_nat) + + +subsection\Variations on Mathematical Induction\ + +(*complete induction*) + +lemmas complete_induct = Ord_induct [OF _ Ord_nat, case_names less, consumes 1] + +lemmas complete_induct_rule = + complete_induct [rule_format, case_names less, consumes 1] + + +lemma nat_induct_from_lemma [rule_format]: + "[| n \ nat; m \ nat; + !!x. [| x \ nat; m \ x; P(x) |] ==> P(succ(x)) |] + ==> m \ n \ P(m) \ P(n)" +apply (erule nat_induct) +apply (simp_all add: distrib_simps le0_iff le_succ_iff) +done + +(*Induction starting from m rather than 0*) +lemma nat_induct_from: + "[| m \ n; m \ nat; n \ nat; + P(m); + !!x. [| x \ nat; m \ x; P(x) |] ==> P(succ(x)) |] + ==> P(n)" +apply (blast intro: nat_induct_from_lemma) +done + +(*Induction suitable for subtraction and less-than*) +lemma diff_induct [case_names 0 0_succ succ_succ, consumes 2]: + "[| m \ nat; n \ nat; + !!x. x \ nat ==> P(x,0); + !!y. y \ nat ==> P(0,succ(y)); + !!x y. [| x \ nat; y \ nat; P(x,y) |] ==> P(succ(x),succ(y)) |] + ==> P(m,n)" +apply (erule_tac x = m in rev_bspec) +apply (erule nat_induct, simp) +apply (rule ballI) +apply (rename_tac i j) +apply (erule_tac n=j in nat_induct, auto) +done + + +(** Induction principle analogous to trancl_induct **) + +lemma succ_lt_induct_lemma [rule_format]: + "m \ nat ==> P(m,succ(m)) \ (\x\nat. P(m,x) \ P(m,succ(x))) \ + (\n\nat. m P(m,n))" +apply (erule nat_induct) + apply (intro impI, rule nat_induct [THEN ballI]) + prefer 4 apply (intro impI, rule nat_induct [THEN ballI]) +apply (auto simp add: le_iff) +done + +lemma succ_lt_induct: + "[| m nat; + P(m,succ(m)); + !!x. [| x \ nat; P(m,x) |] ==> P(m,succ(x)) |] + ==> P(m,n)" +by (blast intro: succ_lt_induct_lemma lt_nat_in_nat) + +subsection\quasinat: to allow a case-split rule for @{term nat_case}\ + +text\True if the argument is zero or any successor\ +lemma [iff]: "quasinat(0)" +by (simp add: quasinat_def) + +lemma [iff]: "quasinat(succ(x))" +by (simp add: quasinat_def) + +lemma nat_imp_quasinat: "n \ nat ==> quasinat(n)" +by (erule natE, simp_all) + +lemma non_nat_case: "~ quasinat(x) ==> nat_case(a,b,x) = 0" +by (simp add: quasinat_def nat_case_def) + +lemma nat_cases_disj: "k=0 | (\y. k = succ(y)) | ~ quasinat(k)" +apply (case_tac "k=0", simp) +apply (case_tac "\m. k = succ(m)") +apply (simp_all add: quasinat_def) +done + +lemma nat_cases: + "[|k=0 ==> P; !!y. k = succ(y) ==> P; ~ quasinat(k) ==> P|] ==> P" +by (insert nat_cases_disj [of k], blast) + +(** nat_case **) + +lemma nat_case_0 [simp]: "nat_case(a,b,0) = a" +by (simp add: nat_case_def) + +lemma nat_case_succ [simp]: "nat_case(a,b,succ(n)) = b(n)" +by (simp add: nat_case_def) + +lemma nat_case_type [TC]: + "[| n \ nat; a \ C(0); !!m. m \ nat ==> b(m): C(succ(m)) |] + ==> nat_case(a,b,n) \ C(n)" +by (erule nat_induct, auto) + +lemma split_nat_case: + "P(nat_case(a,b,k)) \ + ((k=0 \ P(a)) & (\x. k=succ(x) \ P(b(x))) & (~ quasinat(k) \ P(0)))" +apply (rule nat_cases [of k]) +apply (auto simp add: non_nat_case) +done + + +subsection\Recursion on the Natural Numbers\ + +(** nat_rec is used to define eclose and transrec, then becomes obsolete. + The operator rec, from arith.thy, has fewer typing conditions **) + +lemma nat_rec_0: "nat_rec(0,a,b) = a" +apply (rule nat_rec_def [THEN def_wfrec, THEN trans]) + apply (rule wf_Memrel) +apply (rule nat_case_0) +done + +lemma nat_rec_succ: "m \ nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))" +apply (rule nat_rec_def [THEN def_wfrec, THEN trans]) + apply (rule wf_Memrel) +apply (simp add: vimage_singleton_iff) +done + +(** The union of two natural numbers is a natural number -- their maximum **) + +lemma Un_nat_type [TC]: "[| i \ nat; j \ nat |] ==> i \ j \ nat" +apply (rule Un_least_lt [THEN ltD]) +apply (simp_all add: lt_def) +done + +lemma Int_nat_type [TC]: "[| i \ nat; j \ nat |] ==> i \ j \ nat" +apply (rule Int_greatest_lt [THEN ltD]) +apply (simp_all add: lt_def) +done + +(*needed to simplify unions over nat*) +lemma nat_nonempty [simp]: "nat \ 0" +by blast + +text\A natural number is the set of its predecessors\ +lemma nat_eq_Collect_lt: "i \ nat ==> {j\nat. j \ Le \ x \ y & x \ nat & y \ nat" +by (force simp add: Le_def) + +end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Nat_ZF.thy --- a/src/ZF/Nat_ZF.thy Sun Jun 24 11:41:41 2018 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,302 +0,0 @@ -(* Title: ZF/Nat_ZF.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1994 University of Cambridge -*) - -section\The Natural numbers As a Least Fixed Point\ - -theory Nat_ZF imports OrdQuant Bool begin - -definition - nat :: i where - "nat == lfp(Inf, %X. {0} \ {succ(i). i \ X})" - -definition - quasinat :: "i => o" where - "quasinat(n) == n=0 | (\m. n = succ(m))" - -definition - (*Has an unconditional succ case, which is used in "recursor" below.*) - nat_case :: "[i, i=>i, i]=>i" where - "nat_case(a,b,k) == THE y. k=0 & y=a | (\x. k=succ(x) & y=b(x))" - -definition - nat_rec :: "[i, i, [i,i]=>i]=>i" where - "nat_rec(k,a,b) == - wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))" - - (*Internalized relations on the naturals*) - -definition - Le :: i where - "Le == {:nat*nat. x \ y}" - -definition - Lt :: i where - "Lt == {:nat*nat. x < y}" - -definition - Ge :: i where - "Ge == {:nat*nat. y \ x}" - -definition - Gt :: i where - "Gt == {:nat*nat. y < x}" - -definition - greater_than :: "i=>i" where - "greater_than(n) == {i \ nat. n < i}" - -text\No need for a less-than operator: a natural number is its list of -predecessors!\ - - -lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} \ {succ(i). i \ X})" -apply (rule bnd_monoI) -apply (cut_tac infinity, blast, blast) -done - -(* @{term"nat = {0} \ {succ(x). x \ nat}"} *) -lemmas nat_unfold = nat_bnd_mono [THEN nat_def [THEN def_lfp_unfold]] - -(** Type checking of 0 and successor **) - -lemma nat_0I [iff,TC]: "0 \ nat" -apply (subst nat_unfold) -apply (rule singletonI [THEN UnI1]) -done - -lemma nat_succI [intro!,TC]: "n \ nat ==> succ(n) \ nat" -apply (subst nat_unfold) -apply (erule RepFunI [THEN UnI2]) -done - -lemma nat_1I [iff,TC]: "1 \ nat" -by (rule nat_0I [THEN nat_succI]) - -lemma nat_2I [iff,TC]: "2 \ nat" -by (rule nat_1I [THEN nat_succI]) - -lemma bool_subset_nat: "bool \ nat" -by (blast elim!: boolE) - -lemmas bool_into_nat = bool_subset_nat [THEN subsetD] - - -subsection\Injectivity Properties and Induction\ - -(*Mathematical induction*) -lemma nat_induct [case_names 0 succ, induct set: nat]: - "[| n \ nat; P(0); !!x. [| x \ nat; P(x) |] ==> P(succ(x)) |] ==> P(n)" -by (erule def_induct [OF nat_def nat_bnd_mono], blast) - -lemma natE: - assumes "n \ nat" - obtains ("0") "n=0" | (succ) x where "x \ nat" "n=succ(x)" -using assms -by (rule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE]) auto - -lemma nat_into_Ord [simp]: "n \ nat ==> Ord(n)" -by (erule nat_induct, auto) - -(* @{term"i \ nat ==> 0 \ i"}; same thing as @{term"0 nat ==> i \ i"}; same thing as @{term"i nat ==> ~ Limit(a)" -by (induct a rule: nat_induct, auto) - -lemma succ_natD: "succ(i): nat ==> i \ nat" -by (rule Ord_trans [OF succI1], auto) - -lemma nat_succ_iff [iff]: "succ(n): nat \ n \ nat" -by (blast dest!: succ_natD) - -lemma nat_le_Limit: "Limit(i) ==> nat \ i" -apply (rule subset_imp_le) -apply (simp_all add: Limit_is_Ord) -apply (rule subsetI) -apply (erule nat_induct) - apply (erule Limit_has_0 [THEN ltD]) -apply (blast intro: Limit_has_succ [THEN ltD] ltI Limit_is_Ord) -done - -(* [| succ(i): k; k \ nat |] ==> i \ k *) -lemmas succ_in_naturalD = Ord_trans [OF succI1 _ nat_into_Ord] - -lemma lt_nat_in_nat: "[| m nat |] ==> m \ nat" -apply (erule ltE) -apply (erule Ord_trans, assumption, simp) -done - -lemma le_in_nat: "[| m \ n; n \ nat |] ==> m \ nat" -by (blast dest!: lt_nat_in_nat) - - -subsection\Variations on Mathematical Induction\ - -(*complete induction*) - -lemmas complete_induct = Ord_induct [OF _ Ord_nat, case_names less, consumes 1] - -lemmas complete_induct_rule = - complete_induct [rule_format, case_names less, consumes 1] - - -lemma nat_induct_from_lemma [rule_format]: - "[| n \ nat; m \ nat; - !!x. [| x \ nat; m \ x; P(x) |] ==> P(succ(x)) |] - ==> m \ n \ P(m) \ P(n)" -apply (erule nat_induct) -apply (simp_all add: distrib_simps le0_iff le_succ_iff) -done - -(*Induction starting from m rather than 0*) -lemma nat_induct_from: - "[| m \ n; m \ nat; n \ nat; - P(m); - !!x. [| x \ nat; m \ x; P(x) |] ==> P(succ(x)) |] - ==> P(n)" -apply (blast intro: nat_induct_from_lemma) -done - -(*Induction suitable for subtraction and less-than*) -lemma diff_induct [case_names 0 0_succ succ_succ, consumes 2]: - "[| m \ nat; n \ nat; - !!x. x \ nat ==> P(x,0); - !!y. y \ nat ==> P(0,succ(y)); - !!x y. [| x \ nat; y \ nat; P(x,y) |] ==> P(succ(x),succ(y)) |] - ==> P(m,n)" -apply (erule_tac x = m in rev_bspec) -apply (erule nat_induct, simp) -apply (rule ballI) -apply (rename_tac i j) -apply (erule_tac n=j in nat_induct, auto) -done - - -(** Induction principle analogous to trancl_induct **) - -lemma succ_lt_induct_lemma [rule_format]: - "m \ nat ==> P(m,succ(m)) \ (\x\nat. P(m,x) \ P(m,succ(x))) \ - (\n\nat. m P(m,n))" -apply (erule nat_induct) - apply (intro impI, rule nat_induct [THEN ballI]) - prefer 4 apply (intro impI, rule nat_induct [THEN ballI]) -apply (auto simp add: le_iff) -done - -lemma succ_lt_induct: - "[| m nat; - P(m,succ(m)); - !!x. [| x \ nat; P(m,x) |] ==> P(m,succ(x)) |] - ==> P(m,n)" -by (blast intro: succ_lt_induct_lemma lt_nat_in_nat) - -subsection\quasinat: to allow a case-split rule for @{term nat_case}\ - -text\True if the argument is zero or any successor\ -lemma [iff]: "quasinat(0)" -by (simp add: quasinat_def) - -lemma [iff]: "quasinat(succ(x))" -by (simp add: quasinat_def) - -lemma nat_imp_quasinat: "n \ nat ==> quasinat(n)" -by (erule natE, simp_all) - -lemma non_nat_case: "~ quasinat(x) ==> nat_case(a,b,x) = 0" -by (simp add: quasinat_def nat_case_def) - -lemma nat_cases_disj: "k=0 | (\y. k = succ(y)) | ~ quasinat(k)" -apply (case_tac "k=0", simp) -apply (case_tac "\m. k = succ(m)") -apply (simp_all add: quasinat_def) -done - -lemma nat_cases: - "[|k=0 ==> P; !!y. k = succ(y) ==> P; ~ quasinat(k) ==> P|] ==> P" -by (insert nat_cases_disj [of k], blast) - -(** nat_case **) - -lemma nat_case_0 [simp]: "nat_case(a,b,0) = a" -by (simp add: nat_case_def) - -lemma nat_case_succ [simp]: "nat_case(a,b,succ(n)) = b(n)" -by (simp add: nat_case_def) - -lemma nat_case_type [TC]: - "[| n \ nat; a \ C(0); !!m. m \ nat ==> b(m): C(succ(m)) |] - ==> nat_case(a,b,n) \ C(n)" -by (erule nat_induct, auto) - -lemma split_nat_case: - "P(nat_case(a,b,k)) \ - ((k=0 \ P(a)) & (\x. k=succ(x) \ P(b(x))) & (~ quasinat(k) \ P(0)))" -apply (rule nat_cases [of k]) -apply (auto simp add: non_nat_case) -done - - -subsection\Recursion on the Natural Numbers\ - -(** nat_rec is used to define eclose and transrec, then becomes obsolete. - The operator rec, from arith.thy, has fewer typing conditions **) - -lemma nat_rec_0: "nat_rec(0,a,b) = a" -apply (rule nat_rec_def [THEN def_wfrec, THEN trans]) - apply (rule wf_Memrel) -apply (rule nat_case_0) -done - -lemma nat_rec_succ: "m \ nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))" -apply (rule nat_rec_def [THEN def_wfrec, THEN trans]) - apply (rule wf_Memrel) -apply (simp add: vimage_singleton_iff) -done - -(** The union of two natural numbers is a natural number -- their maximum **) - -lemma Un_nat_type [TC]: "[| i \ nat; j \ nat |] ==> i \ j \ nat" -apply (rule Un_least_lt [THEN ltD]) -apply (simp_all add: lt_def) -done - -lemma Int_nat_type [TC]: "[| i \ nat; j \ nat |] ==> i \ j \ nat" -apply (rule Int_greatest_lt [THEN ltD]) -apply (simp_all add: lt_def) -done - -(*needed to simplify unions over nat*) -lemma nat_nonempty [simp]: "nat \ 0" -by blast - -text\A natural number is the set of its predecessors\ -lemma nat_eq_Collect_lt: "i \ nat ==> {j\nat. j \ Le \ x \ y & x \ nat & y \ nat" -by (force simp add: Le_def) - -end diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/OrderType.thy --- a/src/ZF/OrderType.thy Sun Jun 24 11:41:41 2018 +0100 +++ b/src/ZF/OrderType.thy Sun Jun 24 15:57:48 2018 +0200 @@ -5,7 +5,7 @@ section\Order Types and Ordinal Arithmetic\ -theory OrderType imports OrderArith OrdQuant Nat_ZF begin +theory OrderType imports OrderArith OrdQuant Nat begin text\The order type of a well-ordering is the least ordinal isomorphic to it. Ordinal arithmetic is traditionally defined in terms of order types, as it is diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/ZF.thy --- a/src/ZF/ZF.thy Sun Jun 24 11:41:41 2018 +0100 +++ b/src/ZF/ZF.thy Sun Jun 24 15:57:48 2018 +0200 @@ -1,6 +1,6 @@ section\Main ZF Theory: Everything Except AC\ -theory ZF imports List_ZF IntDiv_ZF CardinalArith begin +theory ZF imports List IntDiv CardinalArith begin (*The theory of "iterates" logically belongs to Nat, but can't go there because primrec isn't available into after Datatype.*) diff -r 56034bd1b5f6 -r eb53f944c8cd src/ZF/Zorn.thy --- a/src/ZF/Zorn.thy Sun Jun 24 11:41:41 2018 +0100 +++ b/src/ZF/Zorn.thy Sun Jun 24 15:57:48 2018 +0200 @@ -5,7 +5,7 @@ section\Zorn's Lemma\ -theory Zorn imports OrderArith AC Inductive_ZF begin +theory Zorn imports OrderArith AC Inductive begin text\Based upon the unpublished article ``Towards the Mechanization of the Proofs of Some Classical Theorems of Set Theory,'' by Abrial and Laffitte.\