author | krauss |
Mon, 11 Feb 2008 15:40:21 +0100 | |
changeset 26056 | 6a0801279f4c |
parent 26055 | a7a537e0413a |
child 26057 | f5d5c4922cdf |
--- a/src/ZF/AC.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/AC.thy Mon Feb 11 15:40:21 2008 +0100 @@ -7,7 +7,7 @@ header{*The Axiom of Choice*} -theory AC imports Main begin +theory AC imports Main_ZF begin text{*This definition comes from Halmos (1960), page 59.*} axiomatization where
--- a/src/ZF/Bin.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Bin.thy Mon Feb 11 15:40:21 2008 +0100 @@ -17,7 +17,7 @@ header{*Arithmetic on Binary Integers*} theory Bin -imports Int Datatype +imports Int_ZF Datatype_ZF uses "Tools/numeral_syntax.ML" begin
--- a/src/ZF/Cardinal.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Cardinal.thy Mon Feb 11 15:40:21 2008 +0100 @@ -7,7 +7,7 @@ header{*Cardinal Numbers Without the Axiom of Choice*} -theory Cardinal imports OrderType Finite Nat Sum begin +theory Cardinal imports OrderType Finite Nat_ZF Sum begin definition (*least ordinal operator*)
--- a/src/ZF/Datatype.thy Mon Feb 11 15:19:17 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,114 +0,0 @@ -(* Title: ZF/Datatype.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1997 University of Cambridge - -*) - -header{*Datatype and CoDatatype Definitions*} - -theory Datatype -imports Inductive Univ QUniv -uses "Tools/datatype_package.ML" -begin - -ML_setup {* -(*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 = ref false; - - fun mk_new ([],[]) = Const("True",FOLogic.oT) - | mk_new (largs,rargs) = - BalancedTree.make FOLogic.mk_conj - (map FOLogic.mk_eq (ListPair.zip (largs,rargs))); - - val datatype_ss = @{simpset}; - - fun proc sg ss old = - let val _ = if !trace then writeln ("data_free: OLD = " ^ - string_of_cterm (cterm_of sg 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 sg) lname) - handle Option => raise Match; - val rcon_info = the (Symtab.lookup (ConstructorsData.get sg) rname) - handle 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("False",FOLogic.oT) - else raise Match - val _ = if !trace then - writeln ("NEW = " ^ string_of_cterm (Thm.cterm_of sg new)) - else (); - val goal = Logic.mk_equals (old, new) - val thm = Goal.prove (Simplifier.the_context ss) [] [] goal - (fn _ => rtac iff_reflection 1 THEN - simp_tac (Simplifier.inherit_context ss datatype_ss addsimps #free_iffs lcon_info) 1) - handle ERROR msg => - (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Sign.string_of_term sg goal); - raise Match) - in SOME thm end - handle Match => NONE; - - - val conv = Simplifier.simproc @{theory} "data_free" ["(x::i) = y"] proc; - -end; - - -Addsimprocs [DataFree.conv]; -*} - -end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Datatype_ZF.thy Mon Feb 11 15:40:21 2008 +0100 @@ -0,0 +1,114 @@ +(* Title: ZF/Datatype.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1997 University of Cambridge + +*) + +header{*Datatype and CoDatatype Definitions*} + +theory Datatype_ZF +imports Inductive_ZF Univ QUniv +uses "Tools/datatype_package.ML" +begin + +ML_setup {* +(*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 = ref false; + + fun mk_new ([],[]) = Const("True",FOLogic.oT) + | mk_new (largs,rargs) = + BalancedTree.make FOLogic.mk_conj + (map FOLogic.mk_eq (ListPair.zip (largs,rargs))); + + val datatype_ss = @{simpset}; + + fun proc sg ss old = + let val _ = if !trace then writeln ("data_free: OLD = " ^ + string_of_cterm (cterm_of sg 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 sg) lname) + handle Option => raise Match; + val rcon_info = the (Symtab.lookup (ConstructorsData.get sg) rname) + handle 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("False",FOLogic.oT) + else raise Match + val _ = if !trace then + writeln ("NEW = " ^ string_of_cterm (Thm.cterm_of sg new)) + else (); + val goal = Logic.mk_equals (old, new) + val thm = Goal.prove (Simplifier.the_context ss) [] [] goal + (fn _ => rtac iff_reflection 1 THEN + simp_tac (Simplifier.inherit_context ss datatype_ss addsimps #free_iffs lcon_info) 1) + handle ERROR msg => + (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Sign.string_of_term sg goal); + raise Match) + in SOME thm end + handle Match => NONE; + + + val conv = Simplifier.simproc @{theory} "data_free" ["(x::i) = y"] proc; + +end; + + +Addsimprocs [DataFree.conv]; +*} + +end
--- a/src/ZF/Epsilon.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Epsilon.thy Mon Feb 11 15:40:21 2008 +0100 @@ -7,7 +7,7 @@ header{*Epsilon Induction and Recursion*} -theory Epsilon imports Nat begin +theory Epsilon imports Nat_ZF begin definition eclose :: "i=>i" where
--- a/src/ZF/Finite.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Finite.thy Mon Feb 11 15:40:21 2008 +0100 @@ -8,7 +8,7 @@ header{*Finite Powerset Operator and Finite Function Space*} -theory Finite imports Inductive Epsilon Nat begin +theory Finite imports Inductive_ZF Epsilon Nat_ZF begin (*The natural numbers as a datatype*) rep_datatype
--- a/src/ZF/Induct/Primrec.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Induct/Primrec.thy Mon Feb 11 15:40:21 2008 +0100 @@ -31,7 +31,7 @@ definition COMP :: "[i,i]=>i" where - "COMP(g,fs) == \<lambda>l \<in> list(nat). g ` List.map(\<lambda>f. f`l, fs)" + "COMP(g,fs) == \<lambda>l \<in> list(nat). g ` List_ZF.map(\<lambda>f. f`l, fs)" definition PREC :: "[i,i]=>i" where @@ -93,7 +93,7 @@ monos list_mono con_defs SC_def CONSTANT_def PROJ_def COMP_def PREC_def type_intros nat_typechecks list.intros - lam_type list_case_type drop_type List.map_type + lam_type list_case_type drop_type List_ZF.map_type apply_type rec_type
--- a/src/ZF/Induct/Term.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Induct/Term.thy Mon Feb 11 15:40:21 2008 +0100 @@ -235,7 +235,7 @@ \medskip Theorems about @{text term_map}. *} -declare List.map_compose [simp] +declare List_ZF.map_compose [simp] lemma term_map_ident: "t \<in> term(A) ==> term_map(\<lambda>u. u, t) = t" by (induct rule: term_induct_eqn) simp
--- a/src/ZF/Induct/Tree_Forest.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Induct/Tree_Forest.thy Mon Feb 11 15:40:21 2008 +0100 @@ -247,7 +247,7 @@ *} lemma preorder_map: - "z \<in> tree_forest(A) ==> preorder(map(h,z)) = List.map(h, preorder(z))" + "z \<in> tree_forest(A) ==> preorder(map(h,z)) = List_ZF.map(h, preorder(z))" by (induct set: tree_forest) (simp_all add: map_app_distrib) end
--- a/src/ZF/Inductive.thy Mon Feb 11 15:19:17 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,125 +0,0 @@ -(* Title: ZF/Inductive.thy - ID: $Id$ - 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 -*) - -header{*Inductive and Coinductive Definitions*} - -theory Inductive imports Fixedpt QPair - uses - "ind_syntax.ML" - "Tools/cartprod.ML" - "Tools/ind_cases.ML" - "Tools/inductive_package.ML" - "Tools/induct_tacs.ML" - "Tools/primrec_package.ML" begin - -setup IndCases.setup -setup DatatypeTactics.setup - -ML_setup {* -val iT = Ind_Syntax.iT -and oT = FOLogic.oT; - -structure Lfp = - struct - val oper = Const("Fixedpt.lfp", [iT,iT-->iT]--->iT) - val bnd_mono = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT) - 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", [iT, iT-->iT]--->iT) - val pair = Const("Pair", [iT,iT]--->iT) - val split_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(@{const_name sum}, [iT,iT]--->iT) - val inl = Const("Inl", iT-->iT) - val inr = Const("Inr", iT-->iT) - val elim = Const("case", [iT-->iT, iT-->iT, iT]--->iT) - 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("Fixedpt.gfp", [iT,iT-->iT]--->iT) - val bnd_mono = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT) - 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("QPair.QSigma", [iT, iT-->iT]--->iT) - val pair = Const("QPair.QPair", [iT,iT]--->iT) - val split_name = "QPair.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("QPair.op <+>", [iT,iT]--->iT) - val inl = Const("QPair.QInl", iT-->iT) - val inr = Const("QPair.QInr", iT-->iT) - val elim = Const("QPair.qcase", [iT-->iT, iT-->iT, iT]--->iT) - 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
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Inductive_ZF.thy Mon Feb 11 15:40:21 2008 +0100 @@ -0,0 +1,125 @@ +(* Title: ZF/Inductive.thy + ID: $Id$ + 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 +*) + +header{*Inductive and Coinductive Definitions*} + +theory Inductive_ZF imports Fixedpt QPair Nat_ZF + uses + "ind_syntax.ML" + "Tools/cartprod.ML" + "Tools/ind_cases.ML" + "Tools/inductive_package.ML" + "Tools/induct_tacs.ML" + "Tools/primrec_package.ML" begin + +setup IndCases.setup +setup DatatypeTactics.setup + +ML_setup {* +val iT = Ind_Syntax.iT +and oT = FOLogic.oT; + +structure Lfp = + struct + val oper = Const("Fixedpt.lfp", [iT,iT-->iT]--->iT) + val bnd_mono = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT) + 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", [iT, iT-->iT]--->iT) + val pair = Const("Pair", [iT,iT]--->iT) + val split_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(@{const_name sum}, [iT,iT]--->iT) + val inl = Const("Inl", iT-->iT) + val inr = Const("Inr", iT-->iT) + val elim = Const("case", [iT-->iT, iT-->iT, iT]--->iT) + 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("Fixedpt.gfp", [iT,iT-->iT]--->iT) + val bnd_mono = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT) + 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("QPair.QSigma", [iT, iT-->iT]--->iT) + val pair = Const("QPair.QPair", [iT,iT]--->iT) + val split_name = "QPair.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("QPair.op <+>", [iT,iT]--->iT) + val inl = Const("QPair.QInl", iT-->iT) + val inr = Const("QPair.QInr", iT-->iT) + val elim = Const("QPair.qcase", [iT-->iT, iT-->iT, iT]--->iT) + 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
--- a/src/ZF/InfDatatype.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/InfDatatype.thy Mon Feb 11 15:40:21 2008 +0100 @@ -7,7 +7,7 @@ header{*Infinite-Branching Datatype Definitions*} -theory InfDatatype imports Datatype Univ Finite Cardinal_AC begin +theory InfDatatype imports Datatype_ZF Univ Finite Cardinal_AC begin lemmas fun_Limit_VfromE = Limit_VfromE [OF apply_funtype InfCard_csucc [THEN InfCard_is_Limit]]
--- a/src/ZF/Int.thy Mon Feb 11 15:19:17 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,931 +0,0 @@ -(* Title: ZF/Int.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1993 University of Cambridge - -*) - -header{*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). - \<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & 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 `` {<natify(m), 0>}" - -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) == \<Union><x,y>\<in>z. intrel``{<y,x>}" - -definition - zminus :: "i=>i" ("$- _" [80] 80) where - "$- z == raw_zminus (intify(z))" - -definition - znegative :: "i=>o" where - "znegative(z) == \<exists>x y. x<y & y\<in>nat & <x,y>\<in>z" - -definition - iszero :: "i=>o" where - "iszero(z) == z = $# 0" - -definition - raw_nat_of :: "i=>i" where - "raw_nat_of(z) == natify (\<Union><x,y>\<in>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\<in>nat & ((~ znegative(z) & z = $# m) | - (znegative(z) & $- z = $# m))" - -definition - raw_zmult :: "[i,i]=>i" where - (*Cannot use UN<x1,y2> 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) == - \<Union>p1\<in>z1. \<Union>p2\<in>z2. split(%x1 y1. split(%x2 y2. - intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, 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) == - \<Union>z1\<in>z1. \<Union>z2\<in>z2. let <x1,y1>=z1; <x2,y2>=z2 - in intrel``{<x1#+x2, y1#+y2>}" - -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)" - - -notation (xsymbols) - zmult (infixl "$\<times>" 70) and - zle (infixl "$\<le>" 50) --{*less than or equals*} - -notation (HTML output) - zmult (infixl "$\<times>" 70) and - zle (infixl "$\<le>" 50) - - -declare quotientE [elim!] - -subsection{*Proving that @{term intrel} is an equivalence relation*} - -(** Natural deduction for intrel **) - -lemma intrel_iff [simp]: - "<<x1,y1>,<x2,y2>>: intrel <-> - x1\<in>nat & y1\<in>nat & x2\<in>nat & y2\<in>nat & x1#+y2 = x2#+y1" -by (simp add: intrel_def) - -lemma intrelI [intro!]: - "[| x1#+y2 = x2#+y1; x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] - ==> <<x1,y1>,<x2,y2>>: intrel" -by (simp add: intrel_def) - -lemma intrelE [elim!]: - "[| p: intrel; - !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>; x1#+y2 = x2#+y1; - x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>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\<in>nat; n\<in>nat |] ==> intrel `` {<m,n>} : 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\<in>nat; n\<in>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: "(%<x,y>. intrel``{<y,x>}) 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\<in>nat; y\<in>nat |] ==> raw_zminus(intrel``{<x,y>}) = intrel `` {<y,x>}" -apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent]) -done - -lemma zminus: - "[| x\<in>nat; y\<in>nat |] - ==> $- (intrel``{<x,y>}) = intrel `` {<y,x>}" -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\<in>nat; y\<in>nat |] ==> znegative(intrel``{<x,y>}) <-> x<y" -apply (cases "x<y") -apply (auto simp add: znegative_def not_lt_iff_le) -apply (subgoal_tac "y #+ x2 < x #+ y2", force) -apply (rule add_le_lt_mono, auto) -done - -(*No natural number is negative!*) -lemma not_znegative_int_of [iff]: "~ znegative($# n)" -by (simp add: znegative int_of_def) - -lemma znegative_zminus_int_of [simp]: "znegative($- $# succ(n))" -by (simp add: znegative int_of_def zminus natify_succ) - -lemma not_znegative_imp_zero: "~ znegative($- $# n) ==> 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: "(\<lambda>x. (\<lambda>\<langle>x,y\<rangle>. x #- y)(x)) respects intrel" -by (auto simp add: congruent_def split add: nat_diff_split) - -lemma raw_nat_of: - "[| x\<in>nat; y\<in>nat |] ==> raw_nat_of(intrel``{<x,y>}) = 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) \<in> nat" -by (simp add: raw_nat_of_def) - -lemma nat_of_type [iff,TC]: "nat_of(z) \<in> 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)\<in>nat" -apply (simp add: zmagnitude_def) -apply (rule theI2, auto) -done - -lemma not_zneg_int_of: - "[| z: int; ~ znegative(z) |] ==> \<exists>n\<in>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 |] ==> \<exists>n\<in>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 ==> \<exists>n\<in>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) \<in> int") -apply (simp add: int_def) -apply (auto simp add: znegative nat_of_def raw_nat_of - split add: nat_diff_split) -done - - -subsection{*@{term zadd}: addition on int*} - -text{*Congruence Property for Addition*} -lemma zadd_congruent2: - "(%z1 z2. let <x1,y1>=z1; <x2,y2>=z2 - in intrel``{<x1#+x2, y1#+y2>}) - 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\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] - ==> raw_zadd (intrel``{<x1,y1>}, intrel``{<x2,y2>}) = - intrel `` {<x1#+x2, y1#+y2>}" -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\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] - ==> (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) = - intrel `` {<x1#+x2, y1#+y2>}" -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\<in>nat; n le 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``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, 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\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] - ==> raw_zmult(intrel``{<x1,y1>}, intrel``{<x2,y2>}) = - intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}" -by (simp add: raw_zmult_def - UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2]) - -lemma zmult: - "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] - ==> (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) = - intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}" -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$<w | z=w | w$<z" -apply (simp add: int_def zless_def znegative_def zdiff_def, auto) -apply (simp add: zadd zminus image_iff Bex_def) -apply (rule_tac i = "xb#+ya" and j = "xc #+ y" in Ord_linear_lt) -apply (force dest!: spec simp add: add_ac)+ -done - -lemma zless_linear: "z$<w | intify(z)=intify(w) | w$<z" -apply (cut_tac z = " intify (z) " and w = " intify (w) " in zless_linear_lemma) -apply auto -done - -lemma zless_not_refl [iff]: "~ (z$<z)" -by (auto simp add: zless_def znegative_def int_of_def zdiff_def) - -lemma neq_iff_zless: "[| x: 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 |] ==> (\<exists>n\<in>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 add_left_cancel, auto) -done - -lemma zless_imp_succ_zadd: - "w $< z ==> (\<exists>n\<in>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 <-> (\<exists>n\<in>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\<in>nat; n\<in>nat |] ==> ($#m $< $#n) <-> (m<n)" -apply (simp add: less_iff_succ_add zless_iff_succ_zadd int_of_add [symmetric]) -apply (blast intro: sym) -done - -lemma zless_trans_lemma: - "[| x $< y; y $< z; x: 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 (erule add_left_cancel)+ -apply auto -done - -lemma zless_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, standard] - -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: "[| 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: "[| 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: "[| 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 @{text 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 @{text 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, standard] - -(*"v $<= w ==> z$+v $<= z$+w"*) -lemmas zadd_zless_mono2 = zadd_left_cancel_zless [THEN iffD2, standard] - -(*"v $<= w ==> v$+z $<= w$+z"*) -lemmas zadd_zle_mono1 = zadd_right_cancel_zle [THEN iffD2, standard] - -(*"v $<= w ==> z$+v $<= z$+w"*) -lemmas zadd_zle_mono2 = zadd_left_cancel_zle [THEN iffD2, standard] - -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
--- a/src/ZF/IntDiv.thy Mon Feb 11 15:19:17 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1789 +0,0 @@ -(* Title: ZF/IntDiv.thy - ID: $Id$ - 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 a<b then (0,a) - else let val (q,r) = posDivAlg(a, 2*b) - in if 0<=r-b then (2*q+1, r-b) else (2*q, r) - end - - fun negDivAlg (a,b) = - if 0<=a+b then (~1,a+b) - else let val (q,r) = negDivAlg(a, 2*b) - in if 0<=r-b then (2*q+1, r-b) else (2*q, r) - end; - - fun negateSnd (q,r:int) = (q,~r); - - fun divAlg (a,b) = if 0<=a then - if b>0 then posDivAlg (a,b) - else if a=0 then (0,0) - else negateSnd (negDivAlg (~a,~b)) - else - if 0<b then negDivAlg (a,b) - else negateSnd (posDivAlg (~a,~b)); - -*) - -header{*The Division Operators Div and Mod*} - -theory IntDiv imports IntArith OrderArith begin - -definition - quorem :: "[i,i] => o" where - "quorem == %<a,b> <q,r>. - a = b$*q $+ r & - (#0$<b & #0$<=r & r$<b | ~(#0$<b) & b$<r & r $<= #0)" - -definition - adjust :: "[i,i] => i" where - "adjust(b) == %<q,r>. 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, %<a,b>. nat_of (a $- b $+ #1)), - ab, - %<a,b> f. if (a$<b | b$<=#0) then <#0,a> - else adjust(b, f ` <a,#2$*b>))" - - -(*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, %<a,b>. nat_of ($- a $- b)), - ab, - %<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b> - else adjust(b, f ` <a,#2$*b>))" - -(*for the general case b\<noteq>0*) - -definition - negateSnd :: "i => i" where - "negateSnd == %<q,r>. <q, $-r>" - - (*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 == - %<a,b>. if #0 $<= a then - if #0 $<= b then posDivAlg (<a,b>) - else if a=#0 then <#0,#0> - else negateSnd (negDivAlg (<$-a,$-b>)) - else - if #0$<b then negDivAlg (<a,b>) - else negateSnd (posDivAlg (<$-a,$-b>))" - -definition - zdiv :: "[i,i]=>i" (infixl "zdiv" 70) where - "a zdiv b == fst (divAlg (<intify(a), intify(b)>))" - -definition - zmod :: "[i,i]=>i" (infixl "zmod" 70) where - "a zmod b == snd (divAlg (<intify(a), intify(b)>))" - - -(** 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 \<in> 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 \<in> 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 \<in> 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$<j; k \<in> nat |] ==> 0<k --> $#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$<j; #0 $< k |] ==> 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$<j; #0 $< k |] ==> 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 \<in> int; n \<in> 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 <= and =, - but not (yet?) for k*m < n*k. **) - -lemma zmult_zless_lemma: - "[| k \<in> int; m \<in> int; n \<in> int |] - ==> (m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" -apply (case_tac "k = #0") -apply (auto simp add: neq_iff_zless zmult_zless_mono1 zmult_zless_mono1_neg) -apply (auto simp add: not_zless_iff_zle - not_zle_iff_zless [THEN iff_sym, of "m$*k"] - not_zle_iff_zless [THEN iff_sym, of m]) -apply (auto elim: notE - simp add: zless_imp_zle zmult_zle_mono1 zmult_zle_mono1_neg) -done - -lemma zmult_zless_cancel2: - "(m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" -apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)" - in zmult_zless_lemma) -apply auto -done - -lemma zmult_zless_cancel1: - "(k$*m $< k$*n) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" -by (simp add: zmult_commute [of k] zmult_zless_cancel2) - -lemma zmult_zle_cancel2: - "(m$*k $<= 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 \<in> int; n \<in> int |] ==> m=n <-> (m $<= n & n $<= m)" -apply (blast intro: zle_refl zle_anti_sym) -done - -lemma zmult_cancel2_lemma: - "[| k \<in> int; m \<in> int; n \<in> 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 add: ); -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 (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0; - q \<in> int; q' \<in> 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 (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0; - q \<in> int; q' \<in> int; - r \<in> int; r' \<in> 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 @{text "a\<ge>0"} and @{text "b>0"} *} - -lemma adjust_eq [simp]: - "adjust(b, <q,r>) = (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 $\<times> 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 \<in> int; b \<in> int |] ==> - posDivAlg(<a,b>) = - (if a$<b then <#0,a> else adjust(b, posDivAlg (<a, #2$*b>)))" -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 \<in> int; b \<in> int; - ~ (a $< b | b $<= #0) --> P(<a, #2 $* b>) |] ==> P(<a,b>)" - shows "<u,v> \<in> int*int --> P(<u,v>)" -apply (rule_tac a = "<u,v>" in wf_induct) -apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)" - in wf_measure) -apply clarify -apply (rule prem) -apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec) -apply auto -apply (simp add: not_zle_iff_zless posDivAlg_termination) -done - - -lemma posDivAlg_induct [consumes 2]: - assumes u_int: "u \<in> int" - and v_int: "v \<in> int" - and ih: "!!a b. [| a \<in> int; b \<in> int; - ~ (a $< b | b $<= #0) --> P(a, #2 $* b) |] ==> P(a,b)" - shows "P(u,v)" -apply (subgoal_tac "(%<x,y>. P (x,y)) (<u,v>)") -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 integ_of w \<in> int. - then this rewrite can work for ALL constants!!*) -lemma intify_eq_0_iff_zle: "intify(m) = #0 <-> (m $<= #0 & #0 $<= m)" -apply (simp (no_asm) add: int_eq_iff_zle) -done - - -subsection{* Some convenient biconditionals for products of signs *} - -lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j" -apply (drule zmult_zless_mono1) -apply auto -done - -lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j" -apply (drule zmult_zless_mono1_neg) -apply auto -done - -lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0" -apply (drule zmult_zless_mono1_neg) -apply auto -done - -(** Inequality reasoning **) - -lemma int_0_less_lemma: - "[| x \<in> int; y \<in> 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 \<in> int; y \<in> 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 \<in> int; b \<in> int |] ==> posDivAlg(<a,b>) \<in> 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 add: 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 \<in> int; b \<in> int |] - ==> #0 $<= a --> #0 $< b --> quorem (<a,b>, posDivAlg(<a,b>))" -apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) -apply auto - apply (simp_all add: quorem_def) - txt{*base case: a<b*} - 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(<a,b>) = - (if #0 $<= a$+b then <#-1,a$+b> - else adjust(b, negDivAlg (<a, #2$*b>)))" -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 \<in> int; b \<in> int; - ~ (#0 $<= a $+ b | b $<= #0) --> P(<a, #2 $* b>) |] - ==> P(<a,b>)" - shows "<u,v> \<in> int*int --> P(<u,v>)" -apply (rule_tac a = "<u,v>" in wf_induct) -apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)" - in wf_measure) -apply clarify -apply (rule prem) -apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec) -apply auto -apply (simp add: not_zle_iff_zless negDivAlg_termination) -done - -lemma negDivAlg_induct [consumes 2]: - assumes u_int: "u \<in> int" - and v_int: "v \<in> int" - and ih: "!!a b. [| a \<in> int; b \<in> int; - ~ (#0 $<= a $+ b | b $<= #0) --> P(a, #2 $* b) |] - ==> P(a,b)" - shows "P(u,v)" -apply (subgoal_tac " (%<x,y>. P (x,y)) (<u,v>)") -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 \<in> int; b \<in> int |] ==> negDivAlg(<a,b>) \<in> 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 add: 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 \<in> int; b \<in> int |] - ==> a $< #0 --> #0 $< b --> quorem (<a,b>, negDivAlg(<a,b>))" -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 \<noteq> #0; b \<in> int|] ==> quorem (<#0,b>, <#0,#0>)" -by (force simp add: quorem_def neq_iff_zless) - -lemma posDivAlg_zero_divisor: "posDivAlg(<a,#0>) = <#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 (<q,r>) = <q, $-r>" -apply (unfold negateSnd_def) -apply auto -done - -lemma negateSnd_type: "qr \<in> int * int ==> negateSnd (qr) \<in> int * int" -apply (unfold negateSnd_def) -apply auto -done - -lemma quorem_neg: - "[|quorem (<$-a,$-b>, qr); a \<in> int; b \<in> int; qr \<in> int * int|] - ==> quorem (<a,b>, 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 \<noteq> #0; a \<in> int; b \<in> int|] ==> quorem (<a,b>, divAlg(<a,b>))" -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 \<in> int; b \<in> int|] ==> divAlg(<a,b>) \<in> 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" -apply (simp (no_asm) add: zdiv_def) -done - -lemma zdiv_intify2 [simp]: "x zdiv intify(y) = x zdiv y" -apply (simp (no_asm) add: zdiv_def) -done - -lemma zdiv_type [iff,TC]: "z zdiv w \<in> int" -apply (unfold zdiv_def) -apply (blast intro: fst_type divAlg_type) -done - -lemma zmod_intify1 [simp]: "intify(x) zmod y = x zmod y" -apply (simp (no_asm) add: zmod_def) -done - -lemma zmod_intify2 [simp]: "x zmod intify(y) = x zmod y" -apply (simp (no_asm) add: zmod_def) -done - -lemma zmod_type [iff,TC]: "z zmod w \<in> 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" -apply (simp (no_asm) add: zdiv_def divAlg_def posDivAlg_zero_divisor) -done (*NOT for adding to default simpset*) - -lemma DIVISION_BY_ZERO_ZMOD: "a zmod #0 = intify(a)" -apply (simp (no_asm) add: zmod_def divAlg_def posDivAlg_zero_divisor) -done (*NOT for adding to default simpset*) - - - -(** Basic laws about division and remainder **) - -lemma raw_zmod_zdiv_equality: - "[| a \<in> int; b \<in> 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, standard] -and pos_mod_bound = pos_mod [THEN conjunct2, standard] - -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, standard] -and neg_mod_bound = neg_mod [THEN conjunct2, standard] - - -(** proving general properties of zdiv and zmod **) - -lemma quorem_div_mod: - "[|b \<noteq> #0; a \<in> int; b \<in> int |] - ==> quorem (<a,b>, <a zdiv b, a zmod b>)" -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(<a,b>,<q,r>) implies a \<in> int, but it doesn't matter*) -lemma quorem_div: - "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int |] - ==> a zdiv b = q" -by (blast intro: quorem_div_mod [THEN unique_quotient]) - -lemma quorem_mod: - "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int |] - ==> a zmod b = r" -by (blast intro: quorem_div_mod [THEN unique_remainder]) - -lemma zdiv_pos_pos_trivial_raw: - "[| a \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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,a>,<q,r>); a \<in> int; q \<in> int; a \<noteq> #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,a>,<q,r>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0|] ==> r = #0" -apply (frule self_quotient) -apply (auto simp add: quorem_def) -done - -lemma zdiv_self_raw: "[|a \<noteq> #0; a \<in> int|] ==> a zdiv a = #1" -apply (blast intro: quorem_div_mod [THEN self_quotient]) -done - -lemma zdiv_self [simp]: "intify(a) \<noteq> #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 \<in> 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" -apply (simp (no_asm) add: zdiv_def divAlg_def) -done - -lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def) -done - -lemma zmod_zero [simp]: "#0 zmod b = #0" -apply (simp (no_asm) add: zmod_def divAlg_def) -done - -lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def) -done - -lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1" -apply (simp (no_asm_simp) add: zmod_def divAlg_def) -done - -(** a positive, b positive **) - -lemma zdiv_pos_pos: "[| #0 $< a; #0 $<= b |] - ==> a zdiv b = fst (posDivAlg(<intify(a), intify(b)>))" -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(<intify(a), intify(b)>))" -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(<intify(a), intify(b)>))" -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(<intify(a), intify(b)>))" -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)", standard, simp] -declare zdiv_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zdiv_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zdiv_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zmod_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zmod_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zmod_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zmod_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp] -declare posDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp] -declare negDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp] - - -(** 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 \<in> 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 $\<times> 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 \<in> 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 \<in> 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(<b,c>, <q,r>); c \<in> int; c \<noteq> #0 |] - ==> quorem (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)" -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 \<in> int; c \<in> 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 \<in> int; c \<in> 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) \<noteq> #0 ==> (a$*b) zdiv b = intify(a)" -apply (simp (no_asm_simp) add: zdiv_zmult1_eq) -done - -lemma zdiv_zmult_self2 [simp]: "intify(b) \<noteq> #0 ==> (b$*a) zdiv b = intify(a)" -apply (subst zmult_commute , erule zdiv_zmult_self1) -done - -lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0" -apply (simp (no_asm) add: zmod_zmult1_eq) -done - -lemma zmod_zmult_self2 [simp]: "(b$*a) zmod b = #0" -apply (simp (no_asm) add: zmult_commute zmod_zmult1_eq) -done - - -(** proving (a$+b) zdiv c = - a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **) - -lemma zadd1_lemma: - "[| quorem(<a,c>, <aq,ar>); quorem(<b,c>, <bq,br>); - c \<in> int; c \<noteq> #0 |] - ==> quorem (<a$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)" -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 \<in> int; b \<in> int; c \<in> 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 \<in> int; b \<in> int; c \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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) \<noteq> #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) \<noteq> #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,b>, <q,r>); a \<in> int; b \<in> int; b \<noteq> #0; #0 $< c |] - ==> quorem (<a,b$*c>, <q zdiv c, b$*(q zmod c) $+ r>)" -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 \<in> int; b \<in> 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 \<in> int; b \<in> 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) \<noteq> #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) \<noteq> #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) \<noteq> #0; b \<in> 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) \<noteq> #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) \<noteq> #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) \<noteq> #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) \<noteq> #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 \<in> int; c \<in> 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 Int.thy 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 Int.thy 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 -
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/IntDiv_ZF.thy Mon Feb 11 15:40:21 2008 +0100 @@ -0,0 +1,1789 @@ +(* Title: ZF/IntDiv.thy + ID: $Id$ + 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 a<b then (0,a) + else let val (q,r) = posDivAlg(a, 2*b) + in if 0<=r-b then (2*q+1, r-b) else (2*q, r) + end + + fun negDivAlg (a,b) = + if 0<=a+b then (~1,a+b) + else let val (q,r) = negDivAlg(a, 2*b) + in if 0<=r-b then (2*q+1, r-b) else (2*q, r) + end; + + fun negateSnd (q,r:int) = (q,~r); + + fun divAlg (a,b) = if 0<=a then + if b>0 then posDivAlg (a,b) + else if a=0 then (0,0) + else negateSnd (negDivAlg (~a,~b)) + else + if 0<b then negDivAlg (a,b) + else negateSnd (posDivAlg (~a,~b)); + +*) + +header{*The Division Operators Div and Mod*} + +theory IntDiv_ZF imports IntArith OrderArith begin + +definition + quorem :: "[i,i] => o" where + "quorem == %<a,b> <q,r>. + a = b$*q $+ r & + (#0$<b & #0$<=r & r$<b | ~(#0$<b) & b$<r & r $<= #0)" + +definition + adjust :: "[i,i] => i" where + "adjust(b) == %<q,r>. 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, %<a,b>. nat_of (a $- b $+ #1)), + ab, + %<a,b> f. if (a$<b | b$<=#0) then <#0,a> + else adjust(b, f ` <a,#2$*b>))" + + +(*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, %<a,b>. nat_of ($- a $- b)), + ab, + %<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b> + else adjust(b, f ` <a,#2$*b>))" + +(*for the general case b\<noteq>0*) + +definition + negateSnd :: "i => i" where + "negateSnd == %<q,r>. <q, $-r>" + + (*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 == + %<a,b>. if #0 $<= a then + if #0 $<= b then posDivAlg (<a,b>) + else if a=#0 then <#0,#0> + else negateSnd (negDivAlg (<$-a,$-b>)) + else + if #0$<b then negDivAlg (<a,b>) + else negateSnd (posDivAlg (<$-a,$-b>))" + +definition + zdiv :: "[i,i]=>i" (infixl "zdiv" 70) where + "a zdiv b == fst (divAlg (<intify(a), intify(b)>))" + +definition + zmod :: "[i,i]=>i" (infixl "zmod" 70) where + "a zmod b == snd (divAlg (<intify(a), intify(b)>))" + + +(** 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 \<in> 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 \<in> 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 \<in> 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$<j; k \<in> nat |] ==> 0<k --> $#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$<j; #0 $< k |] ==> 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$<j; #0 $< k |] ==> 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 \<in> int; n \<in> 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 <= and =, + but not (yet?) for k*m < n*k. **) + +lemma zmult_zless_lemma: + "[| k \<in> int; m \<in> int; n \<in> int |] + ==> (m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" +apply (case_tac "k = #0") +apply (auto simp add: neq_iff_zless zmult_zless_mono1 zmult_zless_mono1_neg) +apply (auto simp add: not_zless_iff_zle + not_zle_iff_zless [THEN iff_sym, of "m$*k"] + not_zle_iff_zless [THEN iff_sym, of m]) +apply (auto elim: notE + simp add: zless_imp_zle zmult_zle_mono1 zmult_zle_mono1_neg) +done + +lemma zmult_zless_cancel2: + "(m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" +apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)" + in zmult_zless_lemma) +apply auto +done + +lemma zmult_zless_cancel1: + "(k$*m $< k$*n) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" +by (simp add: zmult_commute [of k] zmult_zless_cancel2) + +lemma zmult_zle_cancel2: + "(m$*k $<= 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 \<in> int; n \<in> int |] ==> m=n <-> (m $<= n & n $<= m)" +apply (blast intro: zle_refl zle_anti_sym) +done + +lemma zmult_cancel2_lemma: + "[| k \<in> int; m \<in> int; n \<in> 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 add: ); +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 (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0; + q \<in> int; q' \<in> 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 (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0; + q \<in> int; q' \<in> int; + r \<in> int; r' \<in> 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 @{text "a\<ge>0"} and @{text "b>0"} *} + +lemma adjust_eq [simp]: + "adjust(b, <q,r>) = (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 $\<times> 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 \<in> int; b \<in> int |] ==> + posDivAlg(<a,b>) = + (if a$<b then <#0,a> else adjust(b, posDivAlg (<a, #2$*b>)))" +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 \<in> int; b \<in> int; + ~ (a $< b | b $<= #0) --> P(<a, #2 $* b>) |] ==> P(<a,b>)" + shows "<u,v> \<in> int*int --> P(<u,v>)" +apply (rule_tac a = "<u,v>" in wf_induct) +apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)" + in wf_measure) +apply clarify +apply (rule prem) +apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec) +apply auto +apply (simp add: not_zle_iff_zless posDivAlg_termination) +done + + +lemma posDivAlg_induct [consumes 2]: + assumes u_int: "u \<in> int" + and v_int: "v \<in> int" + and ih: "!!a b. [| a \<in> int; b \<in> int; + ~ (a $< b | b $<= #0) --> P(a, #2 $* b) |] ==> P(a,b)" + shows "P(u,v)" +apply (subgoal_tac "(%<x,y>. P (x,y)) (<u,v>)") +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 integ_of w \<in> int. + then this rewrite can work for ALL constants!!*) +lemma intify_eq_0_iff_zle: "intify(m) = #0 <-> (m $<= #0 & #0 $<= m)" +apply (simp (no_asm) add: int_eq_iff_zle) +done + + +subsection{* Some convenient biconditionals for products of signs *} + +lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j" +apply (drule zmult_zless_mono1) +apply auto +done + +lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j" +apply (drule zmult_zless_mono1_neg) +apply auto +done + +lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0" +apply (drule zmult_zless_mono1_neg) +apply auto +done + +(** Inequality reasoning **) + +lemma int_0_less_lemma: + "[| x \<in> int; y \<in> 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 \<in> int; y \<in> 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 \<in> int; b \<in> int |] ==> posDivAlg(<a,b>) \<in> 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 add: 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 \<in> int; b \<in> int |] + ==> #0 $<= a --> #0 $< b --> quorem (<a,b>, posDivAlg(<a,b>))" +apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) +apply auto + apply (simp_all add: quorem_def) + txt{*base case: a<b*} + 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(<a,b>) = + (if #0 $<= a$+b then <#-1,a$+b> + else adjust(b, negDivAlg (<a, #2$*b>)))" +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 \<in> int; b \<in> int; + ~ (#0 $<= a $+ b | b $<= #0) --> P(<a, #2 $* b>) |] + ==> P(<a,b>)" + shows "<u,v> \<in> int*int --> P(<u,v>)" +apply (rule_tac a = "<u,v>" in wf_induct) +apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)" + in wf_measure) +apply clarify +apply (rule prem) +apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec) +apply auto +apply (simp add: not_zle_iff_zless negDivAlg_termination) +done + +lemma negDivAlg_induct [consumes 2]: + assumes u_int: "u \<in> int" + and v_int: "v \<in> int" + and ih: "!!a b. [| a \<in> int; b \<in> int; + ~ (#0 $<= a $+ b | b $<= #0) --> P(a, #2 $* b) |] + ==> P(a,b)" + shows "P(u,v)" +apply (subgoal_tac " (%<x,y>. P (x,y)) (<u,v>)") +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 \<in> int; b \<in> int |] ==> negDivAlg(<a,b>) \<in> 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 add: 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 \<in> int; b \<in> int |] + ==> a $< #0 --> #0 $< b --> quorem (<a,b>, negDivAlg(<a,b>))" +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 \<noteq> #0; b \<in> int|] ==> quorem (<#0,b>, <#0,#0>)" +by (force simp add: quorem_def neq_iff_zless) + +lemma posDivAlg_zero_divisor: "posDivAlg(<a,#0>) = <#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 (<q,r>) = <q, $-r>" +apply (unfold negateSnd_def) +apply auto +done + +lemma negateSnd_type: "qr \<in> int * int ==> negateSnd (qr) \<in> int * int" +apply (unfold negateSnd_def) +apply auto +done + +lemma quorem_neg: + "[|quorem (<$-a,$-b>, qr); a \<in> int; b \<in> int; qr \<in> int * int|] + ==> quorem (<a,b>, 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 \<noteq> #0; a \<in> int; b \<in> int|] ==> quorem (<a,b>, divAlg(<a,b>))" +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 \<in> int; b \<in> int|] ==> divAlg(<a,b>) \<in> 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" +apply (simp (no_asm) add: zdiv_def) +done + +lemma zdiv_intify2 [simp]: "x zdiv intify(y) = x zdiv y" +apply (simp (no_asm) add: zdiv_def) +done + +lemma zdiv_type [iff,TC]: "z zdiv w \<in> int" +apply (unfold zdiv_def) +apply (blast intro: fst_type divAlg_type) +done + +lemma zmod_intify1 [simp]: "intify(x) zmod y = x zmod y" +apply (simp (no_asm) add: zmod_def) +done + +lemma zmod_intify2 [simp]: "x zmod intify(y) = x zmod y" +apply (simp (no_asm) add: zmod_def) +done + +lemma zmod_type [iff,TC]: "z zmod w \<in> 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" +apply (simp (no_asm) add: zdiv_def divAlg_def posDivAlg_zero_divisor) +done (*NOT for adding to default simpset*) + +lemma DIVISION_BY_ZERO_ZMOD: "a zmod #0 = intify(a)" +apply (simp (no_asm) add: zmod_def divAlg_def posDivAlg_zero_divisor) +done (*NOT for adding to default simpset*) + + + +(** Basic laws about division and remainder **) + +lemma raw_zmod_zdiv_equality: + "[| a \<in> int; b \<in> 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, standard] +and pos_mod_bound = pos_mod [THEN conjunct2, standard] + +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, standard] +and neg_mod_bound = neg_mod [THEN conjunct2, standard] + + +(** proving general properties of zdiv and zmod **) + +lemma quorem_div_mod: + "[|b \<noteq> #0; a \<in> int; b \<in> int |] + ==> quorem (<a,b>, <a zdiv b, a zmod b>)" +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(<a,b>,<q,r>) implies a \<in> int, but it doesn't matter*) +lemma quorem_div: + "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int |] + ==> a zdiv b = q" +by (blast intro: quorem_div_mod [THEN unique_quotient]) + +lemma quorem_mod: + "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int |] + ==> a zmod b = r" +by (blast intro: quorem_div_mod [THEN unique_remainder]) + +lemma zdiv_pos_pos_trivial_raw: + "[| a \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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,a>,<q,r>); a \<in> int; q \<in> int; a \<noteq> #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,a>,<q,r>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0|] ==> r = #0" +apply (frule self_quotient) +apply (auto simp add: quorem_def) +done + +lemma zdiv_self_raw: "[|a \<noteq> #0; a \<in> int|] ==> a zdiv a = #1" +apply (blast intro: quorem_div_mod [THEN self_quotient]) +done + +lemma zdiv_self [simp]: "intify(a) \<noteq> #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 \<in> 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" +apply (simp (no_asm) add: zdiv_def divAlg_def) +done + +lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def) +done + +lemma zmod_zero [simp]: "#0 zmod b = #0" +apply (simp (no_asm) add: zmod_def divAlg_def) +done + +lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def) +done + +lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1" +apply (simp (no_asm_simp) add: zmod_def divAlg_def) +done + +(** a positive, b positive **) + +lemma zdiv_pos_pos: "[| #0 $< a; #0 $<= b |] + ==> a zdiv b = fst (posDivAlg(<intify(a), intify(b)>))" +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(<intify(a), intify(b)>))" +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(<intify(a), intify(b)>))" +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(<intify(a), intify(b)>))" +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)", standard, simp] +declare zdiv_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zdiv_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zdiv_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zmod_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zmod_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zmod_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zmod_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp] +declare posDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp] +declare negDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp] + + +(** 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 \<in> 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 $\<times> 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 \<in> 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 \<in> 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(<b,c>, <q,r>); c \<in> int; c \<noteq> #0 |] + ==> quorem (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)" +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 \<in> int; c \<in> 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 \<in> int; c \<in> 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) \<noteq> #0 ==> (a$*b) zdiv b = intify(a)" +apply (simp (no_asm_simp) add: zdiv_zmult1_eq) +done + +lemma zdiv_zmult_self2 [simp]: "intify(b) \<noteq> #0 ==> (b$*a) zdiv b = intify(a)" +apply (subst zmult_commute , erule zdiv_zmult_self1) +done + +lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0" +apply (simp (no_asm) add: zmod_zmult1_eq) +done + +lemma zmod_zmult_self2 [simp]: "(b$*a) zmod b = #0" +apply (simp (no_asm) add: zmult_commute zmod_zmult1_eq) +done + + +(** proving (a$+b) zdiv c = + a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **) + +lemma zadd1_lemma: + "[| quorem(<a,c>, <aq,ar>); quorem(<b,c>, <bq,br>); + c \<in> int; c \<noteq> #0 |] + ==> quorem (<a$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)" +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 \<in> int; b \<in> int; c \<in> 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 \<in> int; b \<in> int; c \<in> 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 \<in> int; b \<in> 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 \<in> int; b \<in> 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) \<noteq> #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) \<noteq> #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,b>, <q,r>); a \<in> int; b \<in> int; b \<noteq> #0; #0 $< c |] + ==> quorem (<a,b$*c>, <q zdiv c, b$*(q zmod c) $+ r>)" +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 \<in> int; b \<in> 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 \<in> int; b \<in> 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) \<noteq> #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) \<noteq> #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) \<noteq> #0; b \<in> 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) \<noteq> #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) \<noteq> #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) \<noteq> #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) \<noteq> #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 \<in> int; c \<in> 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 Int.thy 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 Int.thy 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 +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Int_ZF.thy Mon Feb 11 15:40:21 2008 +0100 @@ -0,0 +1,931 @@ +(* Title: ZF/Int.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +*) + +header{*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). + \<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & 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 `` {<natify(m), 0>}" + +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) == \<Union><x,y>\<in>z. intrel``{<y,x>}" + +definition + zminus :: "i=>i" ("$- _" [80] 80) where + "$- z == raw_zminus (intify(z))" + +definition + znegative :: "i=>o" where + "znegative(z) == \<exists>x y. x<y & y\<in>nat & <x,y>\<in>z" + +definition + iszero :: "i=>o" where + "iszero(z) == z = $# 0" + +definition + raw_nat_of :: "i=>i" where + "raw_nat_of(z) == natify (\<Union><x,y>\<in>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\<in>nat & ((~ znegative(z) & z = $# m) | + (znegative(z) & $- z = $# m))" + +definition + raw_zmult :: "[i,i]=>i" where + (*Cannot use UN<x1,y2> 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) == + \<Union>p1\<in>z1. \<Union>p2\<in>z2. split(%x1 y1. split(%x2 y2. + intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, 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) == + \<Union>z1\<in>z1. \<Union>z2\<in>z2. let <x1,y1>=z1; <x2,y2>=z2 + in intrel``{<x1#+x2, y1#+y2>}" + +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)" + + +notation (xsymbols) + zmult (infixl "$\<times>" 70) and + zle (infixl "$\<le>" 50) --{*less than or equals*} + +notation (HTML output) + zmult (infixl "$\<times>" 70) and + zle (infixl "$\<le>" 50) + + +declare quotientE [elim!] + +subsection{*Proving that @{term intrel} is an equivalence relation*} + +(** Natural deduction for intrel **) + +lemma intrel_iff [simp]: + "<<x1,y1>,<x2,y2>>: intrel <-> + x1\<in>nat & y1\<in>nat & x2\<in>nat & y2\<in>nat & x1#+y2 = x2#+y1" +by (simp add: intrel_def) + +lemma intrelI [intro!]: + "[| x1#+y2 = x2#+y1; x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] + ==> <<x1,y1>,<x2,y2>>: intrel" +by (simp add: intrel_def) + +lemma intrelE [elim!]: + "[| p: intrel; + !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>; x1#+y2 = x2#+y1; + x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>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\<in>nat; n\<in>nat |] ==> intrel `` {<m,n>} : 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\<in>nat; n\<in>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: "(%<x,y>. intrel``{<y,x>}) 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\<in>nat; y\<in>nat |] ==> raw_zminus(intrel``{<x,y>}) = intrel `` {<y,x>}" +apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent]) +done + +lemma zminus: + "[| x\<in>nat; y\<in>nat |] + ==> $- (intrel``{<x,y>}) = intrel `` {<y,x>}" +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\<in>nat; y\<in>nat |] ==> znegative(intrel``{<x,y>}) <-> x<y" +apply (cases "x<y") +apply (auto simp add: znegative_def not_lt_iff_le) +apply (subgoal_tac "y #+ x2 < x #+ y2", force) +apply (rule add_le_lt_mono, auto) +done + +(*No natural number is negative!*) +lemma not_znegative_int_of [iff]: "~ znegative($# n)" +by (simp add: znegative int_of_def) + +lemma znegative_zminus_int_of [simp]: "znegative($- $# succ(n))" +by (simp add: znegative int_of_def zminus natify_succ) + +lemma not_znegative_imp_zero: "~ znegative($- $# n) ==> 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: "(\<lambda>x. (\<lambda>\<langle>x,y\<rangle>. x #- y)(x)) respects intrel" +by (auto simp add: congruent_def split add: nat_diff_split) + +lemma raw_nat_of: + "[| x\<in>nat; y\<in>nat |] ==> raw_nat_of(intrel``{<x,y>}) = 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) \<in> nat" +by (simp add: raw_nat_of_def) + +lemma nat_of_type [iff,TC]: "nat_of(z) \<in> 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)\<in>nat" +apply (simp add: zmagnitude_def) +apply (rule theI2, auto) +done + +lemma not_zneg_int_of: + "[| z: int; ~ znegative(z) |] ==> \<exists>n\<in>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 |] ==> \<exists>n\<in>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 ==> \<exists>n\<in>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) \<in> int") +apply (simp add: int_def) +apply (auto simp add: znegative nat_of_def raw_nat_of + split add: nat_diff_split) +done + + +subsection{*@{term zadd}: addition on int*} + +text{*Congruence Property for Addition*} +lemma zadd_congruent2: + "(%z1 z2. let <x1,y1>=z1; <x2,y2>=z2 + in intrel``{<x1#+x2, y1#+y2>}) + 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\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] + ==> raw_zadd (intrel``{<x1,y1>}, intrel``{<x2,y2>}) = + intrel `` {<x1#+x2, y1#+y2>}" +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\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] + ==> (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) = + intrel `` {<x1#+x2, y1#+y2>}" +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\<in>nat; n le 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``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, 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\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] + ==> raw_zmult(intrel``{<x1,y1>}, intrel``{<x2,y2>}) = + intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}" +by (simp add: raw_zmult_def + UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2]) + +lemma zmult: + "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] + ==> (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) = + intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}" +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$<w | z=w | w$<z" +apply (simp add: int_def zless_def znegative_def zdiff_def, auto) +apply (simp add: zadd zminus image_iff Bex_def) +apply (rule_tac i = "xb#+ya" and j = "xc #+ y" in Ord_linear_lt) +apply (force dest!: spec simp add: add_ac)+ +done + +lemma zless_linear: "z$<w | intify(z)=intify(w) | w$<z" +apply (cut_tac z = " intify (z) " and w = " intify (w) " in zless_linear_lemma) +apply auto +done + +lemma zless_not_refl [iff]: "~ (z$<z)" +by (auto simp add: zless_def znegative_def int_of_def zdiff_def) + +lemma neq_iff_zless: "[| x: 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 |] ==> (\<exists>n\<in>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 add_left_cancel, auto) +done + +lemma zless_imp_succ_zadd: + "w $< z ==> (\<exists>n\<in>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 <-> (\<exists>n\<in>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\<in>nat; n\<in>nat |] ==> ($#m $< $#n) <-> (m<n)" +apply (simp add: less_iff_succ_add zless_iff_succ_zadd int_of_add [symmetric]) +apply (blast intro: sym) +done + +lemma zless_trans_lemma: + "[| x $< y; y $< z; x: 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 (erule add_left_cancel)+ +apply auto +done + +lemma zless_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, standard] + +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: "[| 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: "[| 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: "[| 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 @{text 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 @{text 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, standard] + +(*"v $<= w ==> z$+v $<= z$+w"*) +lemmas zadd_zless_mono2 = zadd_left_cancel_zless [THEN iffD2, standard] + +(*"v $<= w ==> v$+z $<= w$+z"*) +lemmas zadd_zle_mono1 = zadd_right_cancel_zle [THEN iffD2, standard] + +(*"v $<= w ==> z$+v $<= z$+w"*) +lemmas zadd_zle_mono2 = zadd_left_cancel_zle [THEN iffD2, standard] + +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
--- a/src/ZF/IsaMakefile Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/IsaMakefile Mon Feb 11 15:40:21 2008 +0100 @@ -30,9 +30,9 @@ $(OUT)/ZF$(ML_SUFFIX): $(OUT)/FOL$(ML_SUFFIX) AC.thy Arith.thy \ ArithSimp.thy Bin.thy Bool.thy Cardinal.thy CardinalArith.thy \ - Cardinal_AC.thy Datatype.thy Epsilon.thy EquivClass.thy Finite.thy \ - Fixedpt.thy Inductive.thy InfDatatype.thy Int.thy IntArith.thy \ - IntDiv.thy List.thy Main.thy Main_ZFC.thy Nat.thy OrdQuant.thy \ + Cardinal_AC.thy Datatype_ZF.thy Epsilon.thy EquivClass.thy Finite.thy \ + Fixedpt.thy Inductive_ZF.thy InfDatatype.thy Int_ZF.thy IntArith.thy \ + IntDiv_ZF.thy List_ZF.thy Main.thy Main_ZF.thy Main_ZFC.thy Nat_ZF.thy OrdQuant.thy \ Order.thy OrderArith.thy OrderType.thy Ordinal.thy Perm.thy QPair.thy \ QUniv.thy ROOT.ML Sum.thy Tools/cartprod.ML Tools/datatype_package.ML \ Tools/ind_cases.ML Tools/induct_tacs.ML Tools/inductive_package.ML \
--- a/src/ZF/List.thy Mon Feb 11 15:19:17 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1251 +0,0 @@ -(* Title: ZF/List - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1994 University of Cambridge - -*) - -header{*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 - "[]" :: i ("[]") - "@List" :: "is => i" ("[(_)]") - -translations - "[x, xs]" == "Cons(x, [xs])" - "[x]" == "Cons(x, [])" - "[]" == "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(lam n:nat. [], - %a l r. lam 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(lam n:nat. 0, - %a l r. lam 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(lam n:nat. Nil, - %u us vs. lam 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 le j then upt(i, j)@[j] else Nil)" - -definition - min :: "[i,i] =>i" where - "min(x, y) == (if x le y then x else y)" - -definition - max :: "[i, i] =>i" where - "max(x, y) == (if x le 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) \<in> list(A) <-> a \<in> A & l \<in> 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 \<in> list(A)|] ==> x \<in> 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) Un 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 **) - -(* c : list(Collect(B,P)) ==> c : list(B) *) -lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD, standard] - -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) ==> - \<forall>x. EX 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) ==> \<forall>i \<in> length(l). (EX 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: - "\<And>l. [| l \<in> list(A); - \<forall>l' \<in> list(A). length(l') < length(l) --> P(l')|] - ==> P(l)" - shows "n \<in> nat ==> \<forall>l \<in> list(A). length(l) < n --> P(l)" -apply (induct_tac n, simp) -apply (blast intro: ih elim!: leE) -done - -theorem list_complete_induct: - "[| l \<in> list(A); - \<And>l. [| l \<in> list(A); - \<forall>l' \<in> 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<min(j,k) <-> i<j & i<k" -apply (unfold min_def) -apply (auto dest!: not_lt_imp_le intro: lt_trans2 lt_trans) -done - -lemma min_succ_succ [simp]: - "[| 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)) le 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<length(xs) <-> xs ~= Nil" -by (erule list.induct, auto) - -lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) <-> (EX 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) ==> \<forall>ys \<in> 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) ==> - \<forall>ys \<in> list(A). \<forall>us \<in> list(A). \<forall>vs \<in> 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) ==> \<forall>ys \<in> 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) ==> \<forall>ys \<in> 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) --> (\<forall>ys \<in> list(A). \<forall>y \<in> A. xs =ys@[y] -->P)-->P" -by (erule list_append_induct, auto) - - -(** more theorems about drop **) - -lemma length_drop [rule_format,simp]: - "n:nat ==> \<forall>xs \<in> 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 ==> \<forall>xs \<in> list(A). length(xs) le n --> drop(n, xs)=Nil" -apply (erule nat_induct) -apply (auto elim: list.cases) -done - -lemma drop_append [rule_format]: - "n:nat ==> - \<forall>xs \<in> 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 ==> \<forall>xs \<in> list(A). \<forall>n \<in> 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 ==> \<forall>xs \<in> list(A). length(xs) le 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) ==> \<forall>n \<in> 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) ==> - \<forall>ys \<in> list(A). \<forall>n \<in> 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 ==> - \<forall>xs \<in> list(A). \<forall>n \<in> 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) ==> \<forall>n. n < length(xs) --> nth(n,xs) : A" -apply (erule list.induct, simp, clarify) -apply (subgoal_tac "n \<in> nat") - apply (erule natE, auto dest!: le_in_nat) -done - -lemma nth_eq_0 [rule_format]: - "xs:list(A) ==> \<forall>n \<in> nat. length(xs) le n --> nth(n,xs) = 0" -apply (erule list.induct, simp, clarify) -apply (erule natE, auto) -done - -lemma nth_append [rule_format]: - "xs:list(A) ==> - \<forall>n \<in> 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. EX i:nat. i<length(xs) & x = nth(i,xs)}" -apply (induct_tac "xs", simp_all) -apply (rule equalityI, auto) -apply (rule_tac x = 0 in bexI, auto) -apply (erule natE, auto) -done - -(* Other theorems about lists *) - -lemma nth_take_lemma [rule_format]: - "k:nat ==> - \<forall>xs \<in> list(A). (\<forall>ys \<in> list(A). k le length(xs) --> k le length(ys) --> - (\<forall>i \<in> nat. i<k --> 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); - \<forall>i \<in> nat. i < length(xs) --> nth(i,xs) = nth(i,ys) |] - ==> xs = ys" -apply (subgoal_tac "length (xs) le 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); (\<forall>i \<in> nat. take(i, xs) = take(i,ys)) |] - ==> xs = ys" -apply (case_tac "length (xs) le length (ys) ") -apply (drule_tac x = "length (ys) " in bspec) -apply (drule_tac [3] not_lt_imp_le) -apply (subgoal_tac [5] "length (ys) le 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 ==> \<forall>i \<in> nat. \<forall>xs \<in> 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\<in>list(A) - ==> \<forall>i. i < length(xs) --> take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]" -apply (induct_tac "xs", auto) -apply (subgoal_tac "i\<in>nat") -apply (erule natE) -apply (auto simp add: le_in_nat) -done - -lemma take_add [rule_format]: - "[|xs\<in>list(A); j\<in>nat|] - ==> \<forall>i\<in>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\<in>list(A) ==> \<forall>n\<in>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,[]) = - (\<lambda>ys \<in> list(B). list_case([], %y l. [], ys))" - - "zip_aux(B,Cons(x,l)) = - (\<lambda>ys \<in> list(B). - list_case(Nil, %y zs. Cons(<x,y>, 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 \<in> list(A) ==> xs \<in> 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 \<in> list(A)|] - ==> \<forall>ys \<in> 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(<x,y>, 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) ==> \<forall>ys \<in> 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) ==> \<forall>ys \<in> 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) |] ==> - \<forall>xs \<in> 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) |] ==> \<forall>ys \<in> 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) ==> \<forall>xs \<in> 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) ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A). - i < length(xs) --> i < length(ys) --> - nth(i,zip(xs, ys)) = <nth(i,xs),nth(i, 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)) = - {<x, y>:A*B. EX 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 |] ==> \<forall>n \<in> 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) ==> \<forall>i \<in> 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) |] ==> \<forall>i \<in> nat. \<forall>j \<in> 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) ==> - \<forall>i \<in> nat. \<forall>j \<in> 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) ==> \<forall>i \<in> 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) ==> - \<forall>i \<in> 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) ==> - \<forall>i \<in> nat. \<forall>xy \<in> A*B. \<forall>xs \<in> 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) ==> - \<forall>i \<in> 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<j then Cons(i, upt(succ(i), j)) else Nil)" -apply (induct_tac "j", auto) -apply (drule not_lt_imp_le) -apply (auto simp: lt_Ord intro: le_anti_sym) -done - -lemma upt_conv_Nil [simp]: "[| j le 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 le j; j:nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]" -by simp - -lemma upt_conv_Cons: - "[| i<j; j: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 le 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 |] ==> - \<forall>i \<in> nat. i #+ m le 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) ==> - \<forall>n \<in> 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 |] ==> - \<forall>i \<in> 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) ==> ALL 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 \<inter> 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 \<in> nat ==> length(repeat(a,n)) = n" -by (induct_tac n, auto) - -lemma repeat_succ_app: "n \<in> 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 \<in> A; n \<in> nat|] ==> repeat(a,n) \<in> list(A)" -by (induct_tac n, auto) - -end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/List_ZF.thy Mon Feb 11 15:40:21 2008 +0100 @@ -0,0 +1,1251 @@ +(* Title: ZF/List + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +*) + +header{*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 + "[]" :: i ("[]") + "@List" :: "is => i" ("[(_)]") + +translations + "[x, xs]" == "Cons(x, [xs])" + "[x]" == "Cons(x, [])" + "[]" == "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(lam n:nat. [], + %a l r. lam 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(lam n:nat. 0, + %a l r. lam 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(lam n:nat. Nil, + %u us vs. lam 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 le j then upt(i, j)@[j] else Nil)" + +definition + min :: "[i,i] =>i" where + "min(x, y) == (if x le y then x else y)" + +definition + max :: "[i, i] =>i" where + "max(x, y) == (if x le 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) \<in> list(A) <-> a \<in> A & l \<in> 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 \<in> list(A)|] ==> x \<in> 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) Un 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 **) + +(* c : list(Collect(B,P)) ==> c : list(B) *) +lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD, standard] + +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) ==> + \<forall>x. EX 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) ==> \<forall>i \<in> length(l). (EX 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: + "\<And>l. [| l \<in> list(A); + \<forall>l' \<in> list(A). length(l') < length(l) --> P(l')|] + ==> P(l)" + shows "n \<in> nat ==> \<forall>l \<in> list(A). length(l) < n --> P(l)" +apply (induct_tac n, simp) +apply (blast intro: ih elim!: leE) +done + +theorem list_complete_induct: + "[| l \<in> list(A); + \<And>l. [| l \<in> list(A); + \<forall>l' \<in> 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<min(j,k) <-> i<j & i<k" +apply (unfold min_def) +apply (auto dest!: not_lt_imp_le intro: lt_trans2 lt_trans) +done + +lemma min_succ_succ [simp]: + "[| 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)) le 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<length(xs) <-> xs ~= Nil" +by (erule list.induct, auto) + +lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) <-> (EX 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) ==> \<forall>ys \<in> 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) ==> + \<forall>ys \<in> list(A). \<forall>us \<in> list(A). \<forall>vs \<in> 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) ==> \<forall>ys \<in> 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) ==> \<forall>ys \<in> 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) --> (\<forall>ys \<in> list(A). \<forall>y \<in> A. xs =ys@[y] -->P)-->P" +by (erule list_append_induct, auto) + + +(** more theorems about drop **) + +lemma length_drop [rule_format,simp]: + "n:nat ==> \<forall>xs \<in> 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 ==> \<forall>xs \<in> list(A). length(xs) le n --> drop(n, xs)=Nil" +apply (erule nat_induct) +apply (auto elim: list.cases) +done + +lemma drop_append [rule_format]: + "n:nat ==> + \<forall>xs \<in> 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 ==> \<forall>xs \<in> list(A). \<forall>n \<in> 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 ==> \<forall>xs \<in> list(A). length(xs) le 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) ==> \<forall>n \<in> 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) ==> + \<forall>ys \<in> list(A). \<forall>n \<in> 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 ==> + \<forall>xs \<in> list(A). \<forall>n \<in> 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) ==> \<forall>n. n < length(xs) --> nth(n,xs) : A" +apply (erule list.induct, simp, clarify) +apply (subgoal_tac "n \<in> nat") + apply (erule natE, auto dest!: le_in_nat) +done + +lemma nth_eq_0 [rule_format]: + "xs:list(A) ==> \<forall>n \<in> nat. length(xs) le n --> nth(n,xs) = 0" +apply (erule list.induct, simp, clarify) +apply (erule natE, auto) +done + +lemma nth_append [rule_format]: + "xs:list(A) ==> + \<forall>n \<in> 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. EX i:nat. i<length(xs) & x = nth(i,xs)}" +apply (induct_tac "xs", simp_all) +apply (rule equalityI, auto) +apply (rule_tac x = 0 in bexI, auto) +apply (erule natE, auto) +done + +(* Other theorems about lists *) + +lemma nth_take_lemma [rule_format]: + "k:nat ==> + \<forall>xs \<in> list(A). (\<forall>ys \<in> list(A). k le length(xs) --> k le length(ys) --> + (\<forall>i \<in> nat. i<k --> 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); + \<forall>i \<in> nat. i < length(xs) --> nth(i,xs) = nth(i,ys) |] + ==> xs = ys" +apply (subgoal_tac "length (xs) le 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); (\<forall>i \<in> nat. take(i, xs) = take(i,ys)) |] + ==> xs = ys" +apply (case_tac "length (xs) le length (ys) ") +apply (drule_tac x = "length (ys) " in bspec) +apply (drule_tac [3] not_lt_imp_le) +apply (subgoal_tac [5] "length (ys) le 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 ==> \<forall>i \<in> nat. \<forall>xs \<in> 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\<in>list(A) + ==> \<forall>i. i < length(xs) --> take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]" +apply (induct_tac "xs", auto) +apply (subgoal_tac "i\<in>nat") +apply (erule natE) +apply (auto simp add: le_in_nat) +done + +lemma take_add [rule_format]: + "[|xs\<in>list(A); j\<in>nat|] + ==> \<forall>i\<in>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\<in>list(A) ==> \<forall>n\<in>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,[]) = + (\<lambda>ys \<in> list(B). list_case([], %y l. [], ys))" + + "zip_aux(B,Cons(x,l)) = + (\<lambda>ys \<in> list(B). + list_case(Nil, %y zs. Cons(<x,y>, 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 \<in> list(A) ==> xs \<in> 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 \<in> list(A)|] + ==> \<forall>ys \<in> 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(<x,y>, 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) ==> \<forall>ys \<in> 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) ==> \<forall>ys \<in> 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) |] ==> + \<forall>xs \<in> 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) |] ==> \<forall>ys \<in> 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) ==> \<forall>xs \<in> 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) ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A). + i < length(xs) --> i < length(ys) --> + nth(i,zip(xs, ys)) = <nth(i,xs),nth(i, 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)) = + {<x, y>:A*B. EX 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 |] ==> \<forall>n \<in> 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) ==> \<forall>i \<in> 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) |] ==> \<forall>i \<in> nat. \<forall>j \<in> 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) ==> + \<forall>i \<in> nat. \<forall>j \<in> 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) ==> \<forall>i \<in> 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) ==> + \<forall>i \<in> 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) ==> + \<forall>i \<in> nat. \<forall>xy \<in> A*B. \<forall>xs \<in> 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) ==> + \<forall>i \<in> 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<j then Cons(i, upt(succ(i), j)) else Nil)" +apply (induct_tac "j", auto) +apply (drule not_lt_imp_le) +apply (auto simp: lt_Ord intro: le_anti_sym) +done + +lemma upt_conv_Nil [simp]: "[| j le 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 le j; j:nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]" +by simp + +lemma upt_conv_Cons: + "[| i<j; j: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 le 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 |] ==> + \<forall>i \<in> nat. i #+ m le 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) ==> + \<forall>n \<in> 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 |] ==> + \<forall>i \<in> 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) ==> ALL 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 \<inter> 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 \<in> nat ==> length(repeat(a,n)) = n" +by (induct_tac n, auto) + +lemma repeat_succ_app: "n \<in> 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 \<in> A; n \<in> nat|] ==> repeat(a,n) \<in> list(A)" +by (induct_tac n, auto) + +end
--- a/src/ZF/Main.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Main.thy Mon Feb 11 15:40:21 2008 +0100 @@ -1,79 +1,5 @@ -(*$Id$*) - -header{*Theory Main: Everything Except AC*} - -theory Main 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.*) - -subsection{* Iteration of the function @{term F} *} - -consts iterates :: "[i=>i,i,i] => i" ("(_^_ '(_'))" [60,1000,1000] 60) - -primrec - "F^0 (x) = x" - "F^(succ(n)) (x) = F(F^n (x))" - -definition - iterates_omega :: "[i=>i,i] => i" where - "iterates_omega(F,x) == \<Union>n\<in>nat. F^n (x)" - -notation (xsymbols) - iterates_omega ("(_^\<omega> '(_'))" [60,1000] 60) -notation (HTML output) - iterates_omega ("(_^\<omega> '(_'))" [60,1000] 60) - -lemma iterates_triv: - "[| n\<in>nat; F(x) = x |] ==> F^n (x) = x" -by (induct n rule: nat_induct, simp_all) - -lemma iterates_type [TC]: - "[| n:nat; a: A; !!x. x:A ==> F(x) : A |] - ==> F^n (a) : A" -by (induct n rule: nat_induct, simp_all) - -lemma iterates_omega_triv: - "F(x) = x ==> F^\<omega> (x) = x" -by (simp add: iterates_omega_def iterates_triv) - -lemma Ord_iterates [simp]: - "[| n\<in>nat; !!i. Ord(i) ==> Ord(F(i)); Ord(x) |] - ==> Ord(F^n (x))" -by (induct n rule: nat_induct, simp_all) - -lemma iterates_commute: "n \<in> nat ==> F(F^n (x)) = F^n (F(x))" -by (induct_tac n, simp_all) - - -subsection{* Transfinite Recursion *} - -text{*Transfinite recursion for definitions based on the - three cases of ordinals*} - -definition - transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where - "transrec3(k, a, b, c) == - transrec(k, \<lambda>x r. - if x=0 then a - else if Limit(x) then c(x, \<lambda>y\<in>x. r`y) - else b(Arith.pred(x), r ` Arith.pred(x)))" - -lemma transrec3_0 [simp]: "transrec3(0,a,b,c) = a" -by (rule transrec3_def [THEN def_transrec, THEN trans], simp) - -lemma transrec3_succ [simp]: - "transrec3(succ(i),a,b,c) = b(i, transrec3(i,a,b,c))" -by (rule transrec3_def [THEN def_transrec, THEN trans], simp) - -lemma transrec3_Limit: - "Limit(i) ==> - transrec3(i,a,b,c) = c(i, \<lambda>j\<in>i. transrec3(j,a,b,c))" -by (rule transrec3_def [THEN def_transrec, THEN trans], force) - - -ML_setup {* - change_simpset (fn ss => ss setmksimps (map mk_eq o Ord_atomize o gen_all)); -*} +theory Main +imports Main_ZF +begin end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Main_ZF.thy Mon Feb 11 15:40:21 2008 +0100 @@ -0,0 +1,79 @@ +(*$Id$*) + +header{*Theory Main: Everything Except AC*} + +theory Main_ZF imports List_ZF IntDiv_ZF CardinalArith begin + +(*The theory of "iterates" logically belongs to Nat, but can't go there because + primrec isn't available into after Datatype.*) + +subsection{* Iteration of the function @{term F} *} + +consts iterates :: "[i=>i,i,i] => i" ("(_^_ '(_'))" [60,1000,1000] 60) + +primrec + "F^0 (x) = x" + "F^(succ(n)) (x) = F(F^n (x))" + +definition + iterates_omega :: "[i=>i,i] => i" where + "iterates_omega(F,x) == \<Union>n\<in>nat. F^n (x)" + +notation (xsymbols) + iterates_omega ("(_^\<omega> '(_'))" [60,1000] 60) +notation (HTML output) + iterates_omega ("(_^\<omega> '(_'))" [60,1000] 60) + +lemma iterates_triv: + "[| n\<in>nat; F(x) = x |] ==> F^n (x) = x" +by (induct n rule: nat_induct, simp_all) + +lemma iterates_type [TC]: + "[| n:nat; a: A; !!x. x:A ==> F(x) : A |] + ==> F^n (a) : A" +by (induct n rule: nat_induct, simp_all) + +lemma iterates_omega_triv: + "F(x) = x ==> F^\<omega> (x) = x" +by (simp add: iterates_omega_def iterates_triv) + +lemma Ord_iterates [simp]: + "[| n\<in>nat; !!i. Ord(i) ==> Ord(F(i)); Ord(x) |] + ==> Ord(F^n (x))" +by (induct n rule: nat_induct, simp_all) + +lemma iterates_commute: "n \<in> nat ==> F(F^n (x)) = F^n (F(x))" +by (induct_tac n, simp_all) + + +subsection{* Transfinite Recursion *} + +text{*Transfinite recursion for definitions based on the + three cases of ordinals*} + +definition + transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where + "transrec3(k, a, b, c) == + transrec(k, \<lambda>x r. + if x=0 then a + else if Limit(x) then c(x, \<lambda>y\<in>x. r`y) + else b(Arith.pred(x), r ` Arith.pred(x)))" + +lemma transrec3_0 [simp]: "transrec3(0,a,b,c) = a" +by (rule transrec3_def [THEN def_transrec, THEN trans], simp) + +lemma transrec3_succ [simp]: + "transrec3(succ(i),a,b,c) = b(i, transrec3(i,a,b,c))" +by (rule transrec3_def [THEN def_transrec, THEN trans], simp) + +lemma transrec3_Limit: + "Limit(i) ==> + transrec3(i,a,b,c) = c(i, \<lambda>j\<in>i. transrec3(j,a,b,c))" +by (rule transrec3_def [THEN def_transrec, THEN trans], force) + + +ML_setup {* + change_simpset (fn ss => ss setmksimps (map mk_eq o Ord_atomize o gen_all)); +*} + +end
--- a/src/ZF/Main_ZFC.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Main_ZFC.thy Mon Feb 11 15:40:21 2008 +0100 @@ -1,3 +1,3 @@ -theory Main_ZFC imports Main InfDatatype begin +theory Main_ZFC imports Main_ZF InfDatatype begin end
--- a/src/ZF/Nat.thy Mon Feb 11 15:19:17 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,302 +0,0 @@ -(* Title: ZF/Nat.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1994 University of Cambridge - -*) - -header{*The Natural numbers As a Least Fixed Point*} - -theory Nat imports OrdQuant Bool begin - -definition - nat :: i where - "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})" - -definition - quasinat :: "i => o" where - "quasinat(n) == n=0 | (\<exists>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 | (EX 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 == {<x,y>:nat*nat. x le y}" - -definition - Lt :: i where - "Lt == {<x, y>:nat*nat. x < y}" - -definition - Ge :: i where - "Ge == {<x,y>:nat*nat. y le x}" - -definition - Gt :: i where - "Gt == {<x,y>: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} Un {succ(i). i:X})" -apply (rule bnd_monoI) -apply (cut_tac infinity, blast, blast) -done - -(* nat = {0} Un {succ(x). x:nat} *) -lemmas nat_unfold = nat_bnd_mono [THEN nat_def [THEN def_lfp_unfold], standard] - -(** 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, standard] - - -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: - "[| n: nat; n=0 ==> P; !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P" -by (erule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE], auto) - -lemma nat_into_Ord [simp]: "n: nat ==> Ord(n)" -by (erule nat_induct, auto) - -(* i: nat ==> 0 le i; same thing as 0<succ(i) *) -lemmas nat_0_le = nat_into_Ord [THEN Ord_0_le, standard] - -(* i: nat ==> i le i; same thing as i<succ(i) *) -lemmas nat_le_refl = nat_into_Ord [THEN le_refl, standard] - -lemma Ord_nat [iff]: "Ord(nat)" -apply (rule OrdI) -apply (erule_tac [2] nat_into_Ord [THEN Ord_is_Transset]) -apply (unfold Transset_def) -apply (rule ballI) -apply (erule nat_induct, auto) -done - -lemma Limit_nat [iff]: "Limit(nat)" -apply (unfold Limit_def) -apply (safe intro!: ltI Ord_nat) -apply (erule ltD) -done - -lemma naturals_not_limit: "a \<in> 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 le 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<n; n: nat |] ==> m: nat" -apply (erule ltE) -apply (erule Ord_trans, assumption, simp) -done - -lemma le_in_nat: "[| m le 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 le x; P(x) |] ==> P(succ(x)) |] - ==> m le 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 le n; m: nat; n: nat; - P(m); - !!x. [| x: nat; m le 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)) --> (ALL x: nat. P(m,x) --> P(m,succ(x))) --> - (ALL n:nat. m<n --> 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<n; n: 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 \<in> 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 | (\<exists>y. k = succ(y)) | ~ quasinat(k)" -apply (case_tac "k=0", simp) -apply (case_tac "\<exists>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)) & (\<forall>x. k=succ(x) --> P(b(x))) & (~ quasinat(k) \<longrightarrow> 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 Un 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 Int 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 \<in> nat ==> {j\<in>nat. j<i} = i" -apply (rule equalityI) -apply (blast dest: ltD) -apply (auto simp add: Ord_mem_iff_lt) -apply (blast intro: lt_trans) -done - -lemma Le_iff [iff]: "<x,y> : Le <-> x le y & x : nat & y : nat" -by (force simp add: Le_def) - -end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Nat_ZF.thy Mon Feb 11 15:40:21 2008 +0100 @@ -0,0 +1,302 @@ +(* Title: ZF/Nat.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +*) + +header{*The Natural numbers As a Least Fixed Point*} + +theory Nat_ZF imports OrdQuant Bool begin + +definition + nat :: i where + "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})" + +definition + quasinat :: "i => o" where + "quasinat(n) == n=0 | (\<exists>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 | (EX 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 == {<x,y>:nat*nat. x le y}" + +definition + Lt :: i where + "Lt == {<x, y>:nat*nat. x < y}" + +definition + Ge :: i where + "Ge == {<x,y>:nat*nat. y le x}" + +definition + Gt :: i where + "Gt == {<x,y>: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} Un {succ(i). i:X})" +apply (rule bnd_monoI) +apply (cut_tac infinity, blast, blast) +done + +(* nat = {0} Un {succ(x). x:nat} *) +lemmas nat_unfold = nat_bnd_mono [THEN nat_def [THEN def_lfp_unfold], standard] + +(** 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, standard] + + +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: + "[| n: nat; n=0 ==> P; !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P" +by (erule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE], auto) + +lemma nat_into_Ord [simp]: "n: nat ==> Ord(n)" +by (erule nat_induct, auto) + +(* i: nat ==> 0 le i; same thing as 0<succ(i) *) +lemmas nat_0_le = nat_into_Ord [THEN Ord_0_le, standard] + +(* i: nat ==> i le i; same thing as i<succ(i) *) +lemmas nat_le_refl = nat_into_Ord [THEN le_refl, standard] + +lemma Ord_nat [iff]: "Ord(nat)" +apply (rule OrdI) +apply (erule_tac [2] nat_into_Ord [THEN Ord_is_Transset]) +apply (unfold Transset_def) +apply (rule ballI) +apply (erule nat_induct, auto) +done + +lemma Limit_nat [iff]: "Limit(nat)" +apply (unfold Limit_def) +apply (safe intro!: ltI Ord_nat) +apply (erule ltD) +done + +lemma naturals_not_limit: "a \<in> 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 le 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<n; n: nat |] ==> m: nat" +apply (erule ltE) +apply (erule Ord_trans, assumption, simp) +done + +lemma le_in_nat: "[| m le 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 le x; P(x) |] ==> P(succ(x)) |] + ==> m le 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 le n; m: nat; n: nat; + P(m); + !!x. [| x: nat; m le 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)) --> (ALL x: nat. P(m,x) --> P(m,succ(x))) --> + (ALL n:nat. m<n --> 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<n; n: 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 \<in> 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 | (\<exists>y. k = succ(y)) | ~ quasinat(k)" +apply (case_tac "k=0", simp) +apply (case_tac "\<exists>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)) & (\<forall>x. k=succ(x) --> P(b(x))) & (~ quasinat(k) \<longrightarrow> 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 Un 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 Int 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 \<in> nat ==> {j\<in>nat. j<i} = i" +apply (rule equalityI) +apply (blast dest: ltD) +apply (auto simp add: Ord_mem_iff_lt) +apply (blast intro: lt_trans) +done + +lemma Le_iff [iff]: "<x,y> : Le <-> x le y & x : nat & y : nat" +by (force simp add: Le_def) + +end
--- a/src/ZF/OrderType.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/OrderType.thy Mon Feb 11 15:40:21 2008 +0100 @@ -7,7 +7,7 @@ header{*Order Types and Ordinal Arithmetic*} -theory OrderType imports OrderArith OrdQuant Nat begin +theory OrderType imports OrderArith OrdQuant Nat_ZF 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
--- a/src/ZF/ROOT.ML Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/ROOT.ML Mon Feb 11 15:40:21 2008 +0100 @@ -8,5 +8,6 @@ Paulson. *) +use_thy "Main"; use_thy "Main_ZFC";
--- a/src/ZF/Tools/datatype_package.ML Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Tools/datatype_package.ML Mon Feb 11 15:40:21 2008 +0100 @@ -49,7 +49,7 @@ fun add_datatype_i (dom_sum, rec_tms) con_ty_lists (monos, type_intrs, type_elims) thy = let val dummy = (*has essential ancestors?*) - Theory.requires thy "Datatype" "(co)datatype definitions"; + Theory.requires thy "Datatype_ZF" "(co)datatype definitions"; val rec_hds = map head_of rec_tms;
--- a/src/ZF/Tools/inductive_package.ML Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Tools/inductive_package.ML Mon Feb 11 15:40:21 2008 +0100 @@ -64,7 +64,7 @@ fun add_inductive_i verbose (rec_tms, dom_sum) intr_specs (monos, con_defs, type_intrs, type_elims) thy = let - val _ = Theory.requires thy "Inductive" "(co)inductive definitions"; + val _ = Theory.requires thy "Inductive_ZF" "(co)inductive definitions"; val (intr_names, intr_tms) = split_list (map fst intr_specs); val case_names = RuleCases.case_names intr_names;
--- a/src/ZF/UNITY/GenPrefix.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/UNITY/GenPrefix.thy Mon Feb 11 15:40:21 2008 +0100 @@ -12,7 +12,7 @@ header{*Charpentier's Generalized Prefix Relation*} theory GenPrefix -imports Main +imports Main_ZF begin definition (*really belongs in ZF/Trancl*)
--- a/src/ZF/UNITY/Monotonicity.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/UNITY/Monotonicity.thy Mon Feb 11 15:40:21 2008 +0100 @@ -69,7 +69,7 @@ lemma take_mono_left: "[| i le j; xs \<in> list(A); j \<in> nat |] ==> <take(i, xs), take(j, xs)> \<in> prefix(A)" -by (blast intro: Nat.le_in_nat take_mono_left_lemma) +by (blast intro: Nat_ZF.le_in_nat take_mono_left_lemma) lemma take_mono_right: "[| <xs,ys> \<in> prefix(A); i \<in> nat |]
--- a/src/ZF/UNITY/State.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/UNITY/State.thy Mon Feb 11 15:40:21 2008 +0100 @@ -11,7 +11,7 @@ header{*UNITY Program States*} -theory State imports Main begin +theory State imports Main_ZF begin consts var :: i datatype var = Var("i \<in> list(nat)")
--- a/src/ZF/Zorn.thy Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/Zorn.thy Mon Feb 11 15:40:21 2008 +0100 @@ -7,7 +7,7 @@ header{*Zorn's Lemma*} -theory Zorn imports OrderArith AC Inductive begin +theory Zorn imports OrderArith AC Inductive_ZF begin text{*Based upon the unpublished article ``Towards the Mechanization of the Proofs of Some Classical Theorems of Set Theory,'' by Abrial and Laffitte.*}
--- a/src/ZF/ind_syntax.ML Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/ind_syntax.ML Mon Feb 11 15:40:21 2008 +0100 @@ -127,7 +127,7 @@ (fn t => "Datatype set not previously declared as constant: " ^ Sign.string_of_term @{theory IFOL} t); val rec_names = (*nat doesn't have to be added*) - "Nat.nat" :: map (#1 o dest_Const) rec_hds + @{const_name "Nat_ZF.nat"} :: map (#1 o dest_Const) rec_hds val u = if co then quniv else univ val cs = (fold o fold) (fn (_, _, _, prems) => prems |> (fold o fold_aterms) (fn t as Const (a, _) => if a mem_string rec_names then I else insert (op =) t
--- a/src/ZF/int_arith.ML Mon Feb 11 15:19:17 2008 +0100 +++ b/src/ZF/int_arith.ML Mon Feb 11 15:40:21 2008 +0100 @@ -95,12 +95,12 @@ (*Utilities*) -val integ_of_const = Const ("Bin.integ_of", iT --> iT); +val integ_of_const = Const (@{const_name "Bin.integ_of"}, iT --> iT); fun mk_numeral n = integ_of_const $ NumeralSyntax.mk_bin n; (*Decodes a binary INTEGER*) -fun dest_numeral (Const("Bin.integ_of", _) $ w) = +fun dest_numeral (Const(@{const_name "Bin.integ_of"}, _) $ w) = (NumeralSyntax.dest_bin w handle Match => raise TERM("Int_Numeral_Simprocs.dest_numeral:1", [w])) | dest_numeral t = raise TERM("Int_Numeral_Simprocs.dest_numeral:2", [t]); @@ -111,11 +111,11 @@ | find_first_numeral past [] = raise TERM("find_first_numeral", []); val zero = mk_numeral 0; -val mk_plus = FOLogic.mk_binop "Int.zadd"; +val mk_plus = FOLogic.mk_binop @{const_name "Int_ZF.zadd"}; val iT = Ind_Syntax.iT; -val zminus_const = Const ("Int.zminus", iT --> iT); +val zminus_const = Const (@{const_name "Int_ZF.zminus"}, iT --> iT); (*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*) fun mk_sum [] = zero @@ -126,30 +126,30 @@ fun long_mk_sum [] = zero | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts); -val dest_plus = FOLogic.dest_bin "Int.zadd" iT; +val dest_plus = FOLogic.dest_bin @{const_name "Int_ZF.zadd"} iT; (*decompose additions AND subtractions as a sum*) -fun dest_summing (pos, Const ("Int.zadd", _) $ t $ u, ts) = +fun dest_summing (pos, Const (@{const_name "Int_ZF.zadd"}, _) $ t $ u, ts) = dest_summing (pos, t, dest_summing (pos, u, ts)) - | dest_summing (pos, Const ("Int.zdiff", _) $ t $ u, ts) = + | dest_summing (pos, Const (@{const_name "Int_ZF.zdiff"}, _) $ t $ u, ts) = dest_summing (pos, t, dest_summing (not pos, u, ts)) | dest_summing (pos, t, ts) = if pos then t::ts else zminus_const$t :: ts; fun dest_sum t = dest_summing (true, t, []); -val mk_diff = FOLogic.mk_binop "Int.zdiff"; -val dest_diff = FOLogic.dest_bin "Int.zdiff" iT; +val mk_diff = FOLogic.mk_binop @{const_name "Int_ZF.zdiff"}; +val dest_diff = FOLogic.dest_bin @{const_name "Int_ZF.zdiff"} iT; val one = mk_numeral 1; -val mk_times = FOLogic.mk_binop "Int.zmult"; +val mk_times = FOLogic.mk_binop @{const_name "Int_ZF.zmult"}; fun mk_prod [] = one | mk_prod [t] = t | mk_prod (t :: ts) = if t = one then mk_prod ts else mk_times (t, mk_prod ts); -val dest_times = FOLogic.dest_bin "Int.zmult" iT; +val dest_times = FOLogic.dest_bin @{const_name "Int_ZF.zmult"} iT; fun dest_prod t = let val (t,u) = dest_times t @@ -160,7 +160,7 @@ fun mk_coeff (k, t) = mk_times (mk_numeral k, t); (*Express t as a product of (possibly) a numeral with other sorted terms*) -fun dest_coeff sign (Const ("Int.zminus", _) $ t) = dest_coeff (~sign) t +fun dest_coeff sign (Const (@{const_name "Int_ZF.zminus"}, _) $ t) = dest_coeff (~sign) t | dest_coeff sign t = let val ts = sort Term.term_ord (dest_prod t) val (n, ts') = find_first_numeral [] ts @@ -254,8 +254,8 @@ structure LessCancelNumerals = CancelNumeralsFun (open CancelNumeralsCommon val prove_conv = ArithData.prove_conv "intless_cancel_numerals" - val mk_bal = FOLogic.mk_binrel "Int.zless" - val dest_bal = FOLogic.dest_bin "Int.zless" iT + val mk_bal = FOLogic.mk_binrel @{const_name "Int_ZF.zless"} + val dest_bal = FOLogic.dest_bin @{const_name "Int_ZF.zless"} iT val bal_add1 = less_add_iff1 RS iff_trans val bal_add2 = less_add_iff2 RS iff_trans ); @@ -263,8 +263,8 @@ structure LeCancelNumerals = CancelNumeralsFun (open CancelNumeralsCommon val prove_conv = ArithData.prove_conv "intle_cancel_numerals" - val mk_bal = FOLogic.mk_binrel "Int.zle" - val dest_bal = FOLogic.dest_bin "Int.zle" iT + val mk_bal = FOLogic.mk_binrel @{const_name "Int_ZF.zle"} + val dest_bal = FOLogic.dest_bin @{const_name "Int_ZF.zle"} iT val bal_add1 = le_add_iff1 RS iff_trans val bal_add2 = le_add_iff2 RS iff_trans );