Made theory names in ZF disjoint from HOL theory names to allow loading both developments
authorkrauss
Mon, 11 Feb 2008 15:40:21 +0100
changeset 260566a0801279f4c
parent 26055 a7a537e0413a
child 26057 f5d5c4922cdf
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
in a single session (but not merge them).
src/ZF/AC.thy
src/ZF/Bin.thy
src/ZF/Cardinal.thy
src/ZF/Datatype.thy
src/ZF/Datatype_ZF.thy
src/ZF/Epsilon.thy
src/ZF/Finite.thy
src/ZF/Induct/Primrec.thy
src/ZF/Induct/Term.thy
src/ZF/Induct/Tree_Forest.thy
src/ZF/Inductive.thy
src/ZF/Inductive_ZF.thy
src/ZF/InfDatatype.thy
src/ZF/Int.thy
src/ZF/IntDiv.thy
src/ZF/IntDiv_ZF.thy
src/ZF/Int_ZF.thy
src/ZF/IsaMakefile
src/ZF/List.thy
src/ZF/List_ZF.thy
src/ZF/Main.thy
src/ZF/Main_ZF.thy
src/ZF/Main_ZFC.thy
src/ZF/Nat.thy
src/ZF/Nat_ZF.thy
src/ZF/OrderType.thy
src/ZF/ROOT.ML
src/ZF/Tools/datatype_package.ML
src/ZF/Tools/inductive_package.ML
src/ZF/UNITY/GenPrefix.thy
src/ZF/UNITY/Monotonicity.thy
src/ZF/UNITY/State.thy
src/ZF/Zorn.thy
src/ZF/ind_syntax.ML
src/ZF/int_arith.ML
     1.1 --- a/src/ZF/AC.thy	Mon Feb 11 15:19:17 2008 +0100
     1.2 +++ b/src/ZF/AC.thy	Mon Feb 11 15:40:21 2008 +0100
     1.3 @@ -7,7 +7,7 @@
     1.4  
     1.5  header{*The Axiom of Choice*}
     1.6  
     1.7 -theory AC imports Main begin
     1.8 +theory AC imports Main_ZF begin
     1.9  
    1.10  text{*This definition comes from Halmos (1960), page 59.*}
    1.11  axiomatization where
     2.1 --- a/src/ZF/Bin.thy	Mon Feb 11 15:19:17 2008 +0100
     2.2 +++ b/src/ZF/Bin.thy	Mon Feb 11 15:40:21 2008 +0100
     2.3 @@ -17,7 +17,7 @@
     2.4  header{*Arithmetic on Binary Integers*}
     2.5  
     2.6  theory Bin
     2.7 -imports Int Datatype
     2.8 +imports Int_ZF Datatype_ZF
     2.9  uses "Tools/numeral_syntax.ML"
    2.10  begin
    2.11  
     3.1 --- a/src/ZF/Cardinal.thy	Mon Feb 11 15:19:17 2008 +0100
     3.2 +++ b/src/ZF/Cardinal.thy	Mon Feb 11 15:40:21 2008 +0100
     3.3 @@ -7,7 +7,7 @@
     3.4  
     3.5  header{*Cardinal Numbers Without the Axiom of Choice*}
     3.6  
     3.7 -theory Cardinal imports OrderType Finite Nat Sum begin
     3.8 +theory Cardinal imports OrderType Finite Nat_ZF Sum begin
     3.9  
    3.10  definition
    3.11    (*least ordinal operator*)
     4.1 --- a/src/ZF/Datatype.thy	Mon Feb 11 15:19:17 2008 +0100
     4.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
     4.3 @@ -1,114 +0,0 @@
     4.4 -(*  Title:      ZF/Datatype.thy
     4.5 -    ID:         $Id$
     4.6 -    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     4.7 -    Copyright   1997  University of Cambridge
     4.8 -
     4.9 -*)
    4.10 -
    4.11 -header{*Datatype and CoDatatype Definitions*}
    4.12 -
    4.13 -theory Datatype
    4.14 -imports Inductive Univ QUniv
    4.15 -uses "Tools/datatype_package.ML"
    4.16 -begin
    4.17 -
    4.18 -ML_setup {*
    4.19 -(*Typechecking rules for most datatypes involving univ*)
    4.20 -structure Data_Arg =
    4.21 -  struct
    4.22 -  val intrs = 
    4.23 -      [@{thm SigmaI}, @{thm InlI}, @{thm InrI},
    4.24 -       @{thm Pair_in_univ}, @{thm Inl_in_univ}, @{thm Inr_in_univ}, 
    4.25 -       @{thm zero_in_univ}, @{thm A_into_univ}, @{thm nat_into_univ}, @{thm UnCI}];
    4.26 -
    4.27 -
    4.28 -  val elims = [make_elim @{thm InlD}, make_elim @{thm InrD},   (*for mutual recursion*)
    4.29 -               @{thm SigmaE}, @{thm sumE}];                    (*allows * and + in spec*)
    4.30 -  end;
    4.31 -
    4.32 -
    4.33 -structure Data_Package = 
    4.34 -  Add_datatype_def_Fun
    4.35 -   (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP
    4.36 -    and Su=Standard_Sum
    4.37 -    and Ind_Package = Ind_Package
    4.38 -    and Datatype_Arg = Data_Arg
    4.39 -    val coind = false);
    4.40 -
    4.41 -
    4.42 -(*Typechecking rules for most codatatypes involving quniv*)
    4.43 -structure CoData_Arg =
    4.44 -  struct
    4.45 -  val intrs = 
    4.46 -      [@{thm QSigmaI}, @{thm QInlI}, @{thm QInrI},
    4.47 -       @{thm QPair_in_quniv}, @{thm QInl_in_quniv}, @{thm QInr_in_quniv}, 
    4.48 -       @{thm zero_in_quniv}, @{thm A_into_quniv}, @{thm nat_into_quniv}, @{thm UnCI}];
    4.49 -
    4.50 -  val elims = [make_elim @{thm QInlD}, make_elim @{thm QInrD},   (*for mutual recursion*)
    4.51 -               @{thm QSigmaE}, @{thm qsumE}];                    (*allows * and + in spec*)
    4.52 -  end;
    4.53 -
    4.54 -structure CoData_Package = 
    4.55 -  Add_datatype_def_Fun
    4.56 -   (structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP
    4.57 -    and Su=Quine_Sum
    4.58 -    and Ind_Package = CoInd_Package
    4.59 -    and Datatype_Arg = CoData_Arg
    4.60 -    val coind = true);
    4.61 -
    4.62 -
    4.63 -
    4.64 -(*Simproc for freeness reasoning: compare datatype constructors for equality*)
    4.65 -structure DataFree =
    4.66 -struct
    4.67 -  val trace = ref false;
    4.68 -
    4.69 -  fun mk_new ([],[]) = Const("True",FOLogic.oT)
    4.70 -    | mk_new (largs,rargs) =
    4.71 -        BalancedTree.make FOLogic.mk_conj
    4.72 -                 (map FOLogic.mk_eq (ListPair.zip (largs,rargs)));
    4.73 -
    4.74 - val datatype_ss = @{simpset};
    4.75 -
    4.76 - fun proc sg ss old =
    4.77 -   let val _ = if !trace then writeln ("data_free: OLD = " ^ 
    4.78 -                                       string_of_cterm (cterm_of sg old))
    4.79 -               else ()
    4.80 -       val (lhs,rhs) = FOLogic.dest_eq old
    4.81 -       val (lhead, largs) = strip_comb lhs
    4.82 -       and (rhead, rargs) = strip_comb rhs
    4.83 -       val lname = #1 (dest_Const lhead) handle TERM _ => raise Match;
    4.84 -       val rname = #1 (dest_Const rhead) handle TERM _ => raise Match;
    4.85 -       val lcon_info = the (Symtab.lookup (ConstructorsData.get sg) lname)
    4.86 -         handle Option => raise Match;
    4.87 -       val rcon_info = the (Symtab.lookup (ConstructorsData.get sg) rname)
    4.88 -         handle Option => raise Match;
    4.89 -       val new = 
    4.90 -           if #big_rec_name lcon_info = #big_rec_name rcon_info 
    4.91 -               andalso not (null (#free_iffs lcon_info)) then
    4.92 -               if lname = rname then mk_new (largs, rargs)
    4.93 -               else Const("False",FOLogic.oT)
    4.94 -           else raise Match
    4.95 -       val _ = if !trace then 
    4.96 -                 writeln ("NEW = " ^ string_of_cterm (Thm.cterm_of sg new))
    4.97 -               else ();
    4.98 -       val goal = Logic.mk_equals (old, new)
    4.99 -       val thm = Goal.prove (Simplifier.the_context ss) [] [] goal
   4.100 -         (fn _ => rtac iff_reflection 1 THEN
   4.101 -           simp_tac (Simplifier.inherit_context ss datatype_ss addsimps #free_iffs lcon_info) 1)
   4.102 -         handle ERROR msg =>
   4.103 -         (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Sign.string_of_term sg goal);
   4.104 -          raise Match)
   4.105 -   in SOME thm end
   4.106 -   handle Match => NONE;
   4.107 -
   4.108 -
   4.109 - val conv = Simplifier.simproc @{theory} "data_free" ["(x::i) = y"] proc;
   4.110 -
   4.111 -end;
   4.112 -
   4.113 -
   4.114 -Addsimprocs [DataFree.conv];
   4.115 -*}
   4.116 -
   4.117 -end
     5.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     5.2 +++ b/src/ZF/Datatype_ZF.thy	Mon Feb 11 15:40:21 2008 +0100
     5.3 @@ -0,0 +1,114 @@
     5.4 +(*  Title:      ZF/Datatype.thy
     5.5 +    ID:         $Id$
     5.6 +    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     5.7 +    Copyright   1997  University of Cambridge
     5.8 +
     5.9 +*)
    5.10 +
    5.11 +header{*Datatype and CoDatatype Definitions*}
    5.12 +
    5.13 +theory Datatype_ZF
    5.14 +imports Inductive_ZF Univ QUniv
    5.15 +uses "Tools/datatype_package.ML"
    5.16 +begin
    5.17 +
    5.18 +ML_setup {*
    5.19 +(*Typechecking rules for most datatypes involving univ*)
    5.20 +structure Data_Arg =
    5.21 +  struct
    5.22 +  val intrs = 
    5.23 +      [@{thm SigmaI}, @{thm InlI}, @{thm InrI},
    5.24 +       @{thm Pair_in_univ}, @{thm Inl_in_univ}, @{thm Inr_in_univ}, 
    5.25 +       @{thm zero_in_univ}, @{thm A_into_univ}, @{thm nat_into_univ}, @{thm UnCI}];
    5.26 +
    5.27 +
    5.28 +  val elims = [make_elim @{thm InlD}, make_elim @{thm InrD},   (*for mutual recursion*)
    5.29 +               @{thm SigmaE}, @{thm sumE}];                    (*allows * and + in spec*)
    5.30 +  end;
    5.31 +
    5.32 +
    5.33 +structure Data_Package = 
    5.34 +  Add_datatype_def_Fun
    5.35 +   (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP
    5.36 +    and Su=Standard_Sum
    5.37 +    and Ind_Package = Ind_Package
    5.38 +    and Datatype_Arg = Data_Arg
    5.39 +    val coind = false);
    5.40 +
    5.41 +
    5.42 +(*Typechecking rules for most codatatypes involving quniv*)
    5.43 +structure CoData_Arg =
    5.44 +  struct
    5.45 +  val intrs = 
    5.46 +      [@{thm QSigmaI}, @{thm QInlI}, @{thm QInrI},
    5.47 +       @{thm QPair_in_quniv}, @{thm QInl_in_quniv}, @{thm QInr_in_quniv}, 
    5.48 +       @{thm zero_in_quniv}, @{thm A_into_quniv}, @{thm nat_into_quniv}, @{thm UnCI}];
    5.49 +
    5.50 +  val elims = [make_elim @{thm QInlD}, make_elim @{thm QInrD},   (*for mutual recursion*)
    5.51 +               @{thm QSigmaE}, @{thm qsumE}];                    (*allows * and + in spec*)
    5.52 +  end;
    5.53 +
    5.54 +structure CoData_Package = 
    5.55 +  Add_datatype_def_Fun
    5.56 +   (structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP
    5.57 +    and Su=Quine_Sum
    5.58 +    and Ind_Package = CoInd_Package
    5.59 +    and Datatype_Arg = CoData_Arg
    5.60 +    val coind = true);
    5.61 +
    5.62 +
    5.63 +
    5.64 +(*Simproc for freeness reasoning: compare datatype constructors for equality*)
    5.65 +structure DataFree =
    5.66 +struct
    5.67 +  val trace = ref false;
    5.68 +
    5.69 +  fun mk_new ([],[]) = Const("True",FOLogic.oT)
    5.70 +    | mk_new (largs,rargs) =
    5.71 +        BalancedTree.make FOLogic.mk_conj
    5.72 +                 (map FOLogic.mk_eq (ListPair.zip (largs,rargs)));
    5.73 +
    5.74 + val datatype_ss = @{simpset};
    5.75 +
    5.76 + fun proc sg ss old =
    5.77 +   let val _ = if !trace then writeln ("data_free: OLD = " ^ 
    5.78 +                                       string_of_cterm (cterm_of sg old))
    5.79 +               else ()
    5.80 +       val (lhs,rhs) = FOLogic.dest_eq old
    5.81 +       val (lhead, largs) = strip_comb lhs
    5.82 +       and (rhead, rargs) = strip_comb rhs
    5.83 +       val lname = #1 (dest_Const lhead) handle TERM _ => raise Match;
    5.84 +       val rname = #1 (dest_Const rhead) handle TERM _ => raise Match;
    5.85 +       val lcon_info = the (Symtab.lookup (ConstructorsData.get sg) lname)
    5.86 +         handle Option => raise Match;
    5.87 +       val rcon_info = the (Symtab.lookup (ConstructorsData.get sg) rname)
    5.88 +         handle Option => raise Match;
    5.89 +       val new = 
    5.90 +           if #big_rec_name lcon_info = #big_rec_name rcon_info 
    5.91 +               andalso not (null (#free_iffs lcon_info)) then
    5.92 +               if lname = rname then mk_new (largs, rargs)
    5.93 +               else Const("False",FOLogic.oT)
    5.94 +           else raise Match
    5.95 +       val _ = if !trace then 
    5.96 +                 writeln ("NEW = " ^ string_of_cterm (Thm.cterm_of sg new))
    5.97 +               else ();
    5.98 +       val goal = Logic.mk_equals (old, new)
    5.99 +       val thm = Goal.prove (Simplifier.the_context ss) [] [] goal
   5.100 +         (fn _ => rtac iff_reflection 1 THEN
   5.101 +           simp_tac (Simplifier.inherit_context ss datatype_ss addsimps #free_iffs lcon_info) 1)
   5.102 +         handle ERROR msg =>
   5.103 +         (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Sign.string_of_term sg goal);
   5.104 +          raise Match)
   5.105 +   in SOME thm end
   5.106 +   handle Match => NONE;
   5.107 +
   5.108 +
   5.109 + val conv = Simplifier.simproc @{theory} "data_free" ["(x::i) = y"] proc;
   5.110 +
   5.111 +end;
   5.112 +
   5.113 +
   5.114 +Addsimprocs [DataFree.conv];
   5.115 +*}
   5.116 +
   5.117 +end
     6.1 --- a/src/ZF/Epsilon.thy	Mon Feb 11 15:19:17 2008 +0100
     6.2 +++ b/src/ZF/Epsilon.thy	Mon Feb 11 15:40:21 2008 +0100
     6.3 @@ -7,7 +7,7 @@
     6.4  
     6.5  header{*Epsilon Induction and Recursion*}
     6.6  
     6.7 -theory Epsilon imports Nat begin
     6.8 +theory Epsilon imports Nat_ZF begin
     6.9  
    6.10  definition
    6.11    eclose    :: "i=>i"  where
     7.1 --- a/src/ZF/Finite.thy	Mon Feb 11 15:19:17 2008 +0100
     7.2 +++ b/src/ZF/Finite.thy	Mon Feb 11 15:40:21 2008 +0100
     7.3 @@ -8,7 +8,7 @@
     7.4  
     7.5  header{*Finite Powerset Operator and Finite Function Space*}
     7.6  
     7.7 -theory Finite imports Inductive Epsilon Nat begin
     7.8 +theory Finite imports Inductive_ZF Epsilon Nat_ZF begin
     7.9  
    7.10  (*The natural numbers as a datatype*)
    7.11  rep_datatype
     8.1 --- a/src/ZF/Induct/Primrec.thy	Mon Feb 11 15:19:17 2008 +0100
     8.2 +++ b/src/ZF/Induct/Primrec.thy	Mon Feb 11 15:40:21 2008 +0100
     8.3 @@ -31,7 +31,7 @@
     8.4  
     8.5  definition
     8.6    COMP :: "[i,i]=>i"  where
     8.7 -  "COMP(g,fs) == \<lambda>l \<in> list(nat). g ` List.map(\<lambda>f. f`l, fs)"
     8.8 +  "COMP(g,fs) == \<lambda>l \<in> list(nat). g ` List_ZF.map(\<lambda>f. f`l, fs)"
     8.9  
    8.10  definition
    8.11    PREC :: "[i,i]=>i"  where
    8.12 @@ -93,7 +93,7 @@
    8.13    monos list_mono
    8.14    con_defs SC_def CONSTANT_def PROJ_def COMP_def PREC_def
    8.15    type_intros nat_typechecks list.intros
    8.16 -    lam_type list_case_type drop_type List.map_type
    8.17 +    lam_type list_case_type drop_type List_ZF.map_type
    8.18      apply_type rec_type
    8.19  
    8.20  
     9.1 --- a/src/ZF/Induct/Term.thy	Mon Feb 11 15:19:17 2008 +0100
     9.2 +++ b/src/ZF/Induct/Term.thy	Mon Feb 11 15:40:21 2008 +0100
     9.3 @@ -235,7 +235,7 @@
     9.4    \medskip Theorems about @{text term_map}.
     9.5  *}
     9.6  
     9.7 -declare List.map_compose [simp]
     9.8 +declare List_ZF.map_compose [simp]
     9.9  
    9.10  lemma term_map_ident: "t \<in> term(A) ==> term_map(\<lambda>u. u, t) = t"
    9.11    by (induct rule: term_induct_eqn) simp
    10.1 --- a/src/ZF/Induct/Tree_Forest.thy	Mon Feb 11 15:19:17 2008 +0100
    10.2 +++ b/src/ZF/Induct/Tree_Forest.thy	Mon Feb 11 15:40:21 2008 +0100
    10.3 @@ -247,7 +247,7 @@
    10.4  *}
    10.5  
    10.6  lemma preorder_map:
    10.7 -    "z \<in> tree_forest(A) ==> preorder(map(h,z)) = List.map(h, preorder(z))"
    10.8 +    "z \<in> tree_forest(A) ==> preorder(map(h,z)) = List_ZF.map(h, preorder(z))"
    10.9    by (induct set: tree_forest) (simp_all add: map_app_distrib)
   10.10  
   10.11  end
    11.1 --- a/src/ZF/Inductive.thy	Mon Feb 11 15:19:17 2008 +0100
    11.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    11.3 @@ -1,125 +0,0 @@
    11.4 -(*  Title:      ZF/Inductive.thy
    11.5 -    ID:         $Id$
    11.6 -    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    11.7 -    Copyright   1993  University of Cambridge
    11.8 -
    11.9 -Inductive definitions use least fixedpoints with standard products and sums
   11.10 -Coinductive definitions use greatest fixedpoints with Quine products and sums
   11.11 -
   11.12 -Sums are used only for mutual recursion;
   11.13 -Products are used only to derive "streamlined" induction rules for relations
   11.14 -*)
   11.15 -
   11.16 -header{*Inductive and Coinductive Definitions*}
   11.17 -
   11.18 -theory Inductive imports Fixedpt QPair
   11.19 -  uses
   11.20 -    "ind_syntax.ML"
   11.21 -    "Tools/cartprod.ML"
   11.22 -    "Tools/ind_cases.ML"
   11.23 -    "Tools/inductive_package.ML"
   11.24 -    "Tools/induct_tacs.ML"
   11.25 -    "Tools/primrec_package.ML" begin
   11.26 -
   11.27 -setup IndCases.setup
   11.28 -setup DatatypeTactics.setup
   11.29 -
   11.30 -ML_setup {*
   11.31 -val iT = Ind_Syntax.iT
   11.32 -and oT = FOLogic.oT;
   11.33 -
   11.34 -structure Lfp =
   11.35 -  struct
   11.36 -  val oper      = Const("Fixedpt.lfp", [iT,iT-->iT]--->iT)
   11.37 -  val bnd_mono  = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT)
   11.38 -  val bnd_monoI = @{thm bnd_monoI}
   11.39 -  val subs      = @{thm def_lfp_subset}
   11.40 -  val Tarski    = @{thm def_lfp_unfold}
   11.41 -  val induct    = @{thm def_induct}
   11.42 -  end;
   11.43 -
   11.44 -structure Standard_Prod =
   11.45 -  struct
   11.46 -  val sigma     = Const("Sigma", [iT, iT-->iT]--->iT)
   11.47 -  val pair      = Const("Pair", [iT,iT]--->iT)
   11.48 -  val split_name = "split"
   11.49 -  val pair_iff  = @{thm Pair_iff}
   11.50 -  val split_eq  = @{thm split}
   11.51 -  val fsplitI   = @{thm splitI}
   11.52 -  val fsplitD   = @{thm splitD}
   11.53 -  val fsplitE   = @{thm splitE}
   11.54 -  end;
   11.55 -
   11.56 -structure Standard_CP = CartProd_Fun (Standard_Prod);
   11.57 -
   11.58 -structure Standard_Sum =
   11.59 -  struct
   11.60 -  val sum       = Const(@{const_name sum}, [iT,iT]--->iT)
   11.61 -  val inl       = Const("Inl", iT-->iT)
   11.62 -  val inr       = Const("Inr", iT-->iT)
   11.63 -  val elim      = Const("case", [iT-->iT, iT-->iT, iT]--->iT)
   11.64 -  val case_inl  = @{thm case_Inl}
   11.65 -  val case_inr  = @{thm case_Inr}
   11.66 -  val inl_iff   = @{thm Inl_iff}
   11.67 -  val inr_iff   = @{thm Inr_iff}
   11.68 -  val distinct  = @{thm Inl_Inr_iff}
   11.69 -  val distinct' = @{thm Inr_Inl_iff}
   11.70 -  val free_SEs  = Ind_Syntax.mk_free_SEs
   11.71 -            [distinct, distinct', inl_iff, inr_iff, Standard_Prod.pair_iff]
   11.72 -  end;
   11.73 -
   11.74 -
   11.75 -structure Ind_Package =
   11.76 -    Add_inductive_def_Fun
   11.77 -      (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP
   11.78 -       and Su=Standard_Sum val coind = false);
   11.79 -
   11.80 -
   11.81 -structure Gfp =
   11.82 -  struct
   11.83 -  val oper      = Const("Fixedpt.gfp", [iT,iT-->iT]--->iT)
   11.84 -  val bnd_mono  = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT)
   11.85 -  val bnd_monoI = @{thm bnd_monoI}
   11.86 -  val subs      = @{thm def_gfp_subset}
   11.87 -  val Tarski    = @{thm def_gfp_unfold}
   11.88 -  val induct    = @{thm def_Collect_coinduct}
   11.89 -  end;
   11.90 -
   11.91 -structure Quine_Prod =
   11.92 -  struct
   11.93 -  val sigma     = Const("QPair.QSigma", [iT, iT-->iT]--->iT)
   11.94 -  val pair      = Const("QPair.QPair", [iT,iT]--->iT)
   11.95 -  val split_name = "QPair.qsplit"
   11.96 -  val pair_iff  = @{thm QPair_iff}
   11.97 -  val split_eq  = @{thm qsplit}
   11.98 -  val fsplitI   = @{thm qsplitI}
   11.99 -  val fsplitD   = @{thm qsplitD}
  11.100 -  val fsplitE   = @{thm qsplitE}
  11.101 -  end;
  11.102 -
  11.103 -structure Quine_CP = CartProd_Fun (Quine_Prod);
  11.104 -
  11.105 -structure Quine_Sum =
  11.106 -  struct
  11.107 -  val sum       = Const("QPair.op <+>", [iT,iT]--->iT)
  11.108 -  val inl       = Const("QPair.QInl", iT-->iT)
  11.109 -  val inr       = Const("QPair.QInr", iT-->iT)
  11.110 -  val elim      = Const("QPair.qcase", [iT-->iT, iT-->iT, iT]--->iT)
  11.111 -  val case_inl  = @{thm qcase_QInl}
  11.112 -  val case_inr  = @{thm qcase_QInr}
  11.113 -  val inl_iff   = @{thm QInl_iff}
  11.114 -  val inr_iff   = @{thm QInr_iff}
  11.115 -  val distinct  = @{thm QInl_QInr_iff}
  11.116 -  val distinct' = @{thm QInr_QInl_iff}
  11.117 -  val free_SEs  = Ind_Syntax.mk_free_SEs
  11.118 -            [distinct, distinct', inl_iff, inr_iff, Quine_Prod.pair_iff]
  11.119 -  end;
  11.120 -
  11.121 -
  11.122 -structure CoInd_Package =
  11.123 -  Add_inductive_def_Fun(structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP
  11.124 -    and Su=Quine_Sum val coind = true);
  11.125 -
  11.126 -*}
  11.127 -
  11.128 -end
    12.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    12.2 +++ b/src/ZF/Inductive_ZF.thy	Mon Feb 11 15:40:21 2008 +0100
    12.3 @@ -0,0 +1,125 @@
    12.4 +(*  Title:      ZF/Inductive.thy
    12.5 +    ID:         $Id$
    12.6 +    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    12.7 +    Copyright   1993  University of Cambridge
    12.8 +
    12.9 +Inductive definitions use least fixedpoints with standard products and sums
   12.10 +Coinductive definitions use greatest fixedpoints with Quine products and sums
   12.11 +
   12.12 +Sums are used only for mutual recursion;
   12.13 +Products are used only to derive "streamlined" induction rules for relations
   12.14 +*)
   12.15 +
   12.16 +header{*Inductive and Coinductive Definitions*}
   12.17 +
   12.18 +theory Inductive_ZF imports Fixedpt QPair Nat_ZF
   12.19 +  uses
   12.20 +    "ind_syntax.ML"
   12.21 +    "Tools/cartprod.ML"
   12.22 +    "Tools/ind_cases.ML"
   12.23 +    "Tools/inductive_package.ML"
   12.24 +    "Tools/induct_tacs.ML"
   12.25 +    "Tools/primrec_package.ML" begin
   12.26 +
   12.27 +setup IndCases.setup
   12.28 +setup DatatypeTactics.setup
   12.29 +
   12.30 +ML_setup {*
   12.31 +val iT = Ind_Syntax.iT
   12.32 +and oT = FOLogic.oT;
   12.33 +
   12.34 +structure Lfp =
   12.35 +  struct
   12.36 +  val oper      = Const("Fixedpt.lfp", [iT,iT-->iT]--->iT)
   12.37 +  val bnd_mono  = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT)
   12.38 +  val bnd_monoI = @{thm bnd_monoI}
   12.39 +  val subs      = @{thm def_lfp_subset}
   12.40 +  val Tarski    = @{thm def_lfp_unfold}
   12.41 +  val induct    = @{thm def_induct}
   12.42 +  end;
   12.43 +
   12.44 +structure Standard_Prod =
   12.45 +  struct
   12.46 +  val sigma     = Const("Sigma", [iT, iT-->iT]--->iT)
   12.47 +  val pair      = Const("Pair", [iT,iT]--->iT)
   12.48 +  val split_name = "split"
   12.49 +  val pair_iff  = @{thm Pair_iff}
   12.50 +  val split_eq  = @{thm split}
   12.51 +  val fsplitI   = @{thm splitI}
   12.52 +  val fsplitD   = @{thm splitD}
   12.53 +  val fsplitE   = @{thm splitE}
   12.54 +  end;
   12.55 +
   12.56 +structure Standard_CP = CartProd_Fun (Standard_Prod);
   12.57 +
   12.58 +structure Standard_Sum =
   12.59 +  struct
   12.60 +  val sum       = Const(@{const_name sum}, [iT,iT]--->iT)
   12.61 +  val inl       = Const("Inl", iT-->iT)
   12.62 +  val inr       = Const("Inr", iT-->iT)
   12.63 +  val elim      = Const("case", [iT-->iT, iT-->iT, iT]--->iT)
   12.64 +  val case_inl  = @{thm case_Inl}
   12.65 +  val case_inr  = @{thm case_Inr}
   12.66 +  val inl_iff   = @{thm Inl_iff}
   12.67 +  val inr_iff   = @{thm Inr_iff}
   12.68 +  val distinct  = @{thm Inl_Inr_iff}
   12.69 +  val distinct' = @{thm Inr_Inl_iff}
   12.70 +  val free_SEs  = Ind_Syntax.mk_free_SEs
   12.71 +            [distinct, distinct', inl_iff, inr_iff, Standard_Prod.pair_iff]
   12.72 +  end;
   12.73 +
   12.74 +
   12.75 +structure Ind_Package =
   12.76 +    Add_inductive_def_Fun
   12.77 +      (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP
   12.78 +       and Su=Standard_Sum val coind = false);
   12.79 +
   12.80 +
   12.81 +structure Gfp =
   12.82 +  struct
   12.83 +  val oper      = Const("Fixedpt.gfp", [iT,iT-->iT]--->iT)
   12.84 +  val bnd_mono  = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT)
   12.85 +  val bnd_monoI = @{thm bnd_monoI}
   12.86 +  val subs      = @{thm def_gfp_subset}
   12.87 +  val Tarski    = @{thm def_gfp_unfold}
   12.88 +  val induct    = @{thm def_Collect_coinduct}
   12.89 +  end;
   12.90 +
   12.91 +structure Quine_Prod =
   12.92 +  struct
   12.93 +  val sigma     = Const("QPair.QSigma", [iT, iT-->iT]--->iT)
   12.94 +  val pair      = Const("QPair.QPair", [iT,iT]--->iT)
   12.95 +  val split_name = "QPair.qsplit"
   12.96 +  val pair_iff  = @{thm QPair_iff}
   12.97 +  val split_eq  = @{thm qsplit}
   12.98 +  val fsplitI   = @{thm qsplitI}
   12.99 +  val fsplitD   = @{thm qsplitD}
  12.100 +  val fsplitE   = @{thm qsplitE}
  12.101 +  end;
  12.102 +
  12.103 +structure Quine_CP = CartProd_Fun (Quine_Prod);
  12.104 +
  12.105 +structure Quine_Sum =
  12.106 +  struct
  12.107 +  val sum       = Const("QPair.op <+>", [iT,iT]--->iT)
  12.108 +  val inl       = Const("QPair.QInl", iT-->iT)
  12.109 +  val inr       = Const("QPair.QInr", iT-->iT)
  12.110 +  val elim      = Const("QPair.qcase", [iT-->iT, iT-->iT, iT]--->iT)
  12.111 +  val case_inl  = @{thm qcase_QInl}
  12.112 +  val case_inr  = @{thm qcase_QInr}
  12.113 +  val inl_iff   = @{thm QInl_iff}
  12.114 +  val inr_iff   = @{thm QInr_iff}
  12.115 +  val distinct  = @{thm QInl_QInr_iff}
  12.116 +  val distinct' = @{thm QInr_QInl_iff}
  12.117 +  val free_SEs  = Ind_Syntax.mk_free_SEs
  12.118 +            [distinct, distinct', inl_iff, inr_iff, Quine_Prod.pair_iff]
  12.119 +  end;
  12.120 +
  12.121 +
  12.122 +structure CoInd_Package =
  12.123 +  Add_inductive_def_Fun(structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP
  12.124 +    and Su=Quine_Sum val coind = true);
  12.125 +
  12.126 +*}
  12.127 +
  12.128 +end
    13.1 --- a/src/ZF/InfDatatype.thy	Mon Feb 11 15:19:17 2008 +0100
    13.2 +++ b/src/ZF/InfDatatype.thy	Mon Feb 11 15:40:21 2008 +0100
    13.3 @@ -7,7 +7,7 @@
    13.4  
    13.5  header{*Infinite-Branching Datatype Definitions*}
    13.6  
    13.7 -theory InfDatatype imports Datatype Univ Finite Cardinal_AC begin
    13.8 +theory InfDatatype imports Datatype_ZF Univ Finite Cardinal_AC begin
    13.9  
   13.10  lemmas fun_Limit_VfromE = 
   13.11      Limit_VfromE [OF apply_funtype InfCard_csucc [THEN InfCard_is_Limit]]
    14.1 --- a/src/ZF/Int.thy	Mon Feb 11 15:19:17 2008 +0100
    14.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    14.3 @@ -1,931 +0,0 @@
    14.4 -(*  Title:      ZF/Int.thy
    14.5 -    ID:         $Id$
    14.6 -    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    14.7 -    Copyright   1993  University of Cambridge
    14.8 -
    14.9 -*)
   14.10 -
   14.11 -header{*The Integers as Equivalence Classes Over Pairs of Natural Numbers*}
   14.12 -
   14.13 -theory Int imports EquivClass ArithSimp begin
   14.14 -
   14.15 -definition
   14.16 -  intrel :: i  where
   14.17 -    "intrel == {p : (nat*nat)*(nat*nat).                 
   14.18 -                \<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1}"
   14.19 -
   14.20 -definition
   14.21 -  int :: i  where
   14.22 -    "int == (nat*nat)//intrel"  
   14.23 -
   14.24 -definition
   14.25 -  int_of :: "i=>i" --{*coercion from nat to int*}    ("$# _" [80] 80)  where
   14.26 -    "$# m == intrel `` {<natify(m), 0>}"
   14.27 -
   14.28 -definition
   14.29 -  intify :: "i=>i" --{*coercion from ANYTHING to int*}  where
   14.30 -    "intify(m) == if m : int then m else $#0"
   14.31 -
   14.32 -definition
   14.33 -  raw_zminus :: "i=>i"  where
   14.34 -    "raw_zminus(z) == \<Union><x,y>\<in>z. intrel``{<y,x>}"
   14.35 -
   14.36 -definition
   14.37 -  zminus :: "i=>i"                                 ("$- _" [80] 80)  where
   14.38 -    "$- z == raw_zminus (intify(z))"
   14.39 -
   14.40 -definition
   14.41 -  znegative   ::      "i=>o"  where
   14.42 -    "znegative(z) == \<exists>x y. x<y & y\<in>nat & <x,y>\<in>z"
   14.43 -
   14.44 -definition
   14.45 -  iszero      ::      "i=>o"  where
   14.46 -    "iszero(z) == z = $# 0"
   14.47 -    
   14.48 -definition
   14.49 -  raw_nat_of  :: "i=>i"  where
   14.50 -  "raw_nat_of(z) == natify (\<Union><x,y>\<in>z. x#-y)"
   14.51 -
   14.52 -definition
   14.53 -  nat_of  :: "i=>i"  where
   14.54 -  "nat_of(z) == raw_nat_of (intify(z))"
   14.55 -
   14.56 -definition
   14.57 -  zmagnitude  ::      "i=>i"  where
   14.58 -  --{*could be replaced by an absolute value function from int to int?*}
   14.59 -    "zmagnitude(z) ==
   14.60 -     THE m. m\<in>nat & ((~ znegative(z) & z = $# m) |
   14.61 -		       (znegative(z) & $- z = $# m))"
   14.62 -
   14.63 -definition
   14.64 -  raw_zmult   ::      "[i,i]=>i"  where
   14.65 -    (*Cannot use UN<x1,y2> here or in zadd because of the form of congruent2.
   14.66 -      Perhaps a "curried" or even polymorphic congruent predicate would be
   14.67 -      better.*)
   14.68 -     "raw_zmult(z1,z2) == 
   14.69 -       \<Union>p1\<in>z1. \<Union>p2\<in>z2.  split(%x1 y1. split(%x2 y2.        
   14.70 -                   intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)"
   14.71 -
   14.72 -definition
   14.73 -  zmult       ::      "[i,i]=>i"      (infixl "$*" 70)  where
   14.74 -     "z1 $* z2 == raw_zmult (intify(z1),intify(z2))"
   14.75 -
   14.76 -definition
   14.77 -  raw_zadd    ::      "[i,i]=>i"  where
   14.78 -     "raw_zadd (z1, z2) == 
   14.79 -       \<Union>z1\<in>z1. \<Union>z2\<in>z2. let <x1,y1>=z1; <x2,y2>=z2                 
   14.80 -                           in intrel``{<x1#+x2, y1#+y2>}"
   14.81 -
   14.82 -definition
   14.83 -  zadd        ::      "[i,i]=>i"      (infixl "$+" 65)  where
   14.84 -     "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))"
   14.85 -
   14.86 -definition
   14.87 -  zdiff        ::      "[i,i]=>i"      (infixl "$-" 65)  where
   14.88 -     "z1 $- z2 == z1 $+ zminus(z2)"
   14.89 -
   14.90 -definition
   14.91 -  zless        ::      "[i,i]=>o"      (infixl "$<" 50)  where
   14.92 -     "z1 $< z2 == znegative(z1 $- z2)"
   14.93 -  
   14.94 -definition
   14.95 -  zle          ::      "[i,i]=>o"      (infixl "$<=" 50)  where
   14.96 -     "z1 $<= z2 == z1 $< z2 | intify(z1)=intify(z2)"
   14.97 -  
   14.98 -
   14.99 -notation (xsymbols)
  14.100 -  zmult  (infixl "$\<times>" 70) and
  14.101 -  zle  (infixl "$\<le>" 50)  --{*less than or equals*}
  14.102 -
  14.103 -notation (HTML output)
  14.104 -  zmult  (infixl "$\<times>" 70) and
  14.105 -  zle  (infixl "$\<le>" 50)
  14.106 -
  14.107 -
  14.108 -declare quotientE [elim!]
  14.109 -
  14.110 -subsection{*Proving that @{term intrel} is an equivalence relation*}
  14.111 -
  14.112 -(** Natural deduction for intrel **)
  14.113 -
  14.114 -lemma intrel_iff [simp]: 
  14.115 -    "<<x1,y1>,<x2,y2>>: intrel <->  
  14.116 -     x1\<in>nat & y1\<in>nat & x2\<in>nat & y2\<in>nat & x1#+y2 = x2#+y1"
  14.117 -by (simp add: intrel_def)
  14.118 -
  14.119 -lemma intrelI [intro!]: 
  14.120 -    "[| x1#+y2 = x2#+y1; x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |]   
  14.121 -     ==> <<x1,y1>,<x2,y2>>: intrel"
  14.122 -by (simp add: intrel_def)
  14.123 -
  14.124 -lemma intrelE [elim!]:
  14.125 -  "[| p: intrel;   
  14.126 -      !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>;  x1#+y2 = x2#+y1;  
  14.127 -                        x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] ==> Q |]  
  14.128 -   ==> Q"
  14.129 -by (simp add: intrel_def, blast) 
  14.130 -
  14.131 -lemma int_trans_lemma:
  14.132 -     "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2 |] ==> x1 #+ y3 = x3 #+ y1"
  14.133 -apply (rule sym)
  14.134 -apply (erule add_left_cancel)+
  14.135 -apply (simp_all (no_asm_simp))
  14.136 -done
  14.137 -
  14.138 -lemma equiv_intrel: "equiv(nat*nat, intrel)"
  14.139 -apply (simp add: equiv_def refl_def sym_def trans_def)
  14.140 -apply (fast elim!: sym int_trans_lemma)
  14.141 -done
  14.142 -
  14.143 -lemma image_intrel_int: "[| m\<in>nat; n\<in>nat |] ==> intrel `` {<m,n>} : int"
  14.144 -by (simp add: int_def)
  14.145 -
  14.146 -declare equiv_intrel [THEN eq_equiv_class_iff, simp]
  14.147 -declare conj_cong [cong]
  14.148 -
  14.149 -lemmas eq_intrelD = eq_equiv_class [OF _ equiv_intrel]
  14.150 -
  14.151 -(** int_of: the injection from nat to int **)
  14.152 -
  14.153 -lemma int_of_type [simp,TC]: "$#m : int"
  14.154 -by (simp add: int_def quotient_def int_of_def, auto)
  14.155 -
  14.156 -lemma int_of_eq [iff]: "($# m = $# n) <-> natify(m)=natify(n)"
  14.157 -by (simp add: int_of_def)
  14.158 -
  14.159 -lemma int_of_inject: "[| $#m = $#n;  m\<in>nat;  n\<in>nat |] ==> m=n"
  14.160 -by (drule int_of_eq [THEN iffD1], auto)
  14.161 -
  14.162 -
  14.163 -(** intify: coercion from anything to int **)
  14.164 -
  14.165 -lemma intify_in_int [iff,TC]: "intify(x) : int"
  14.166 -by (simp add: intify_def)
  14.167 -
  14.168 -lemma intify_ident [simp]: "n : int ==> intify(n) = n"
  14.169 -by (simp add: intify_def)
  14.170 -
  14.171 -
  14.172 -subsection{*Collapsing rules: to remove @{term intify}
  14.173 -            from arithmetic expressions*}
  14.174 -
  14.175 -lemma intify_idem [simp]: "intify(intify(x)) = intify(x)"
  14.176 -by simp
  14.177 -
  14.178 -lemma int_of_natify [simp]: "$# (natify(m)) = $# m"
  14.179 -by (simp add: int_of_def)
  14.180 -
  14.181 -lemma zminus_intify [simp]: "$- (intify(m)) = $- m"
  14.182 -by (simp add: zminus_def)
  14.183 -
  14.184 -(** Addition **)
  14.185 -
  14.186 -lemma zadd_intify1 [simp]: "intify(x) $+ y = x $+ y"
  14.187 -by (simp add: zadd_def)
  14.188 -
  14.189 -lemma zadd_intify2 [simp]: "x $+ intify(y) = x $+ y"
  14.190 -by (simp add: zadd_def)
  14.191 -
  14.192 -(** Subtraction **)
  14.193 -
  14.194 -lemma zdiff_intify1 [simp]:"intify(x) $- y = x $- y"
  14.195 -by (simp add: zdiff_def)
  14.196 -
  14.197 -lemma zdiff_intify2 [simp]:"x $- intify(y) = x $- y"
  14.198 -by (simp add: zdiff_def)
  14.199 -
  14.200 -(** Multiplication **)
  14.201 -
  14.202 -lemma zmult_intify1 [simp]:"intify(x) $* y = x $* y"
  14.203 -by (simp add: zmult_def)
  14.204 -
  14.205 -lemma zmult_intify2 [simp]:"x $* intify(y) = x $* y"
  14.206 -by (simp add: zmult_def)
  14.207 -
  14.208 -(** Orderings **)
  14.209 -
  14.210 -lemma zless_intify1 [simp]:"intify(x) $< y <-> x $< y"
  14.211 -by (simp add: zless_def)
  14.212 -
  14.213 -lemma zless_intify2 [simp]:"x $< intify(y) <-> x $< y"
  14.214 -by (simp add: zless_def)
  14.215 -
  14.216 -lemma zle_intify1 [simp]:"intify(x) $<= y <-> x $<= y"
  14.217 -by (simp add: zle_def)
  14.218 -
  14.219 -lemma zle_intify2 [simp]:"x $<= intify(y) <-> x $<= y"
  14.220 -by (simp add: zle_def)
  14.221 -
  14.222 -
  14.223 -subsection{*@{term zminus}: unary negation on @{term int}*}
  14.224 -
  14.225 -lemma zminus_congruent: "(%<x,y>. intrel``{<y,x>}) respects intrel"
  14.226 -by (auto simp add: congruent_def add_ac)
  14.227 -
  14.228 -lemma raw_zminus_type: "z : int ==> raw_zminus(z) : int"
  14.229 -apply (simp add: int_def raw_zminus_def)
  14.230 -apply (typecheck add: UN_equiv_class_type [OF equiv_intrel zminus_congruent])
  14.231 -done
  14.232 -
  14.233 -lemma zminus_type [TC,iff]: "$-z : int"
  14.234 -by (simp add: zminus_def raw_zminus_type)
  14.235 -
  14.236 -lemma raw_zminus_inject: 
  14.237 -     "[| raw_zminus(z) = raw_zminus(w);  z: int;  w: int |] ==> z=w"
  14.238 -apply (simp add: int_def raw_zminus_def)
  14.239 -apply (erule UN_equiv_class_inject [OF equiv_intrel zminus_congruent], safe)
  14.240 -apply (auto dest: eq_intrelD simp add: add_ac)
  14.241 -done
  14.242 -
  14.243 -lemma zminus_inject_intify [dest!]: "$-z = $-w ==> intify(z) = intify(w)"
  14.244 -apply (simp add: zminus_def)
  14.245 -apply (blast dest!: raw_zminus_inject)
  14.246 -done
  14.247 -
  14.248 -lemma zminus_inject: "[| $-z = $-w;  z: int;  w: int |] ==> z=w"
  14.249 -by auto
  14.250 -
  14.251 -lemma raw_zminus: 
  14.252 -    "[| x\<in>nat;  y\<in>nat |] ==> raw_zminus(intrel``{<x,y>}) = intrel `` {<y,x>}"
  14.253 -apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent])
  14.254 -done
  14.255 -
  14.256 -lemma zminus: 
  14.257 -    "[| x\<in>nat;  y\<in>nat |]  
  14.258 -     ==> $- (intrel``{<x,y>}) = intrel `` {<y,x>}"
  14.259 -by (simp add: zminus_def raw_zminus image_intrel_int)
  14.260 -
  14.261 -lemma raw_zminus_zminus: "z : int ==> raw_zminus (raw_zminus(z)) = z"
  14.262 -by (auto simp add: int_def raw_zminus)
  14.263 -
  14.264 -lemma zminus_zminus_intify [simp]: "$- ($- z) = intify(z)"
  14.265 -by (simp add: zminus_def raw_zminus_type raw_zminus_zminus)
  14.266 -
  14.267 -lemma zminus_int0 [simp]: "$- ($#0) = $#0"
  14.268 -by (simp add: int_of_def zminus)
  14.269 -
  14.270 -lemma zminus_zminus: "z : int ==> $- ($- z) = z"
  14.271 -by simp
  14.272 -
  14.273 -
  14.274 -subsection{*@{term znegative}: the test for negative integers*}
  14.275 -
  14.276 -lemma znegative: "[| x\<in>nat; y\<in>nat |] ==> znegative(intrel``{<x,y>}) <-> x<y"
  14.277 -apply (cases "x<y") 
  14.278 -apply (auto simp add: znegative_def not_lt_iff_le)
  14.279 -apply (subgoal_tac "y #+ x2 < x #+ y2", force) 
  14.280 -apply (rule add_le_lt_mono, auto) 
  14.281 -done
  14.282 -
  14.283 -(*No natural number is negative!*)
  14.284 -lemma not_znegative_int_of [iff]: "~ znegative($# n)"
  14.285 -by (simp add: znegative int_of_def) 
  14.286 -
  14.287 -lemma znegative_zminus_int_of [simp]: "znegative($- $# succ(n))"
  14.288 -by (simp add: znegative int_of_def zminus natify_succ)
  14.289 -
  14.290 -lemma not_znegative_imp_zero: "~ znegative($- $# n) ==> natify(n)=0"
  14.291 -by (simp add: znegative int_of_def zminus Ord_0_lt_iff [THEN iff_sym])
  14.292 -
  14.293 -
  14.294 -subsection{*@{term nat_of}: Coercion of an Integer to a Natural Number*}
  14.295 -
  14.296 -lemma nat_of_intify [simp]: "nat_of(intify(z)) = nat_of(z)"
  14.297 -by (simp add: nat_of_def)
  14.298 -
  14.299 -lemma nat_of_congruent: "(\<lambda>x. (\<lambda>\<langle>x,y\<rangle>. x #- y)(x)) respects intrel"
  14.300 -by (auto simp add: congruent_def split add: nat_diff_split)
  14.301 -
  14.302 -lemma raw_nat_of: 
  14.303 -    "[| x\<in>nat;  y\<in>nat |] ==> raw_nat_of(intrel``{<x,y>}) = x#-y"
  14.304 -by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent])
  14.305 -
  14.306 -lemma raw_nat_of_int_of: "raw_nat_of($# n) = natify(n)"
  14.307 -by (simp add: int_of_def raw_nat_of)
  14.308 -
  14.309 -lemma nat_of_int_of [simp]: "nat_of($# n) = natify(n)"
  14.310 -by (simp add: raw_nat_of_int_of nat_of_def)
  14.311 -
  14.312 -lemma raw_nat_of_type: "raw_nat_of(z) \<in> nat"
  14.313 -by (simp add: raw_nat_of_def)
  14.314 -
  14.315 -lemma nat_of_type [iff,TC]: "nat_of(z) \<in> nat"
  14.316 -by (simp add: nat_of_def raw_nat_of_type)
  14.317 -
  14.318 -subsection{*zmagnitude: magnitide of an integer, as a natural number*}
  14.319 -
  14.320 -lemma zmagnitude_int_of [simp]: "zmagnitude($# n) = natify(n)"
  14.321 -by (auto simp add: zmagnitude_def int_of_eq)
  14.322 -
  14.323 -lemma natify_int_of_eq: "natify(x)=n ==> $#x = $# n"
  14.324 -apply (drule sym)
  14.325 -apply (simp (no_asm_simp) add: int_of_eq)
  14.326 -done
  14.327 -
  14.328 -lemma zmagnitude_zminus_int_of [simp]: "zmagnitude($- $# n) = natify(n)"
  14.329 -apply (simp add: zmagnitude_def)
  14.330 -apply (rule the_equality)
  14.331 -apply (auto dest!: not_znegative_imp_zero natify_int_of_eq
  14.332 -            iff del: int_of_eq, auto)
  14.333 -done
  14.334 -
  14.335 -lemma zmagnitude_type [iff,TC]: "zmagnitude(z)\<in>nat"
  14.336 -apply (simp add: zmagnitude_def)
  14.337 -apply (rule theI2, auto)
  14.338 -done
  14.339 -
  14.340 -lemma not_zneg_int_of: 
  14.341 -     "[| z: int; ~ znegative(z) |] ==> \<exists>n\<in>nat. z = $# n"
  14.342 -apply (auto simp add: int_def znegative int_of_def not_lt_iff_le)
  14.343 -apply (rename_tac x y) 
  14.344 -apply (rule_tac x="x#-y" in bexI) 
  14.345 -apply (auto simp add: add_diff_inverse2) 
  14.346 -done
  14.347 -
  14.348 -lemma not_zneg_mag [simp]:
  14.349 -     "[| z: int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z"
  14.350 -by (drule not_zneg_int_of, auto)
  14.351 -
  14.352 -lemma zneg_int_of: 
  14.353 -     "[| znegative(z); z: int |] ==> \<exists>n\<in>nat. z = $- ($# succ(n))"
  14.354 -by (auto simp add: int_def znegative zminus int_of_def dest!: less_imp_succ_add)
  14.355 -
  14.356 -lemma zneg_mag [simp]:
  14.357 -     "[| znegative(z); z: int |] ==> $# (zmagnitude(z)) = $- z"
  14.358 -by (drule zneg_int_of, auto)
  14.359 -
  14.360 -lemma int_cases: "z : int ==> \<exists>n\<in>nat. z = $# n | z = $- ($# succ(n))"
  14.361 -apply (case_tac "znegative (z) ")
  14.362 -prefer 2 apply (blast dest: not_zneg_mag sym)
  14.363 -apply (blast dest: zneg_int_of)
  14.364 -done
  14.365 -
  14.366 -lemma not_zneg_raw_nat_of:
  14.367 -     "[| ~ znegative(z); z: int |] ==> $# (raw_nat_of(z)) = z"
  14.368 -apply (drule not_zneg_int_of)
  14.369 -apply (auto simp add: raw_nat_of_type raw_nat_of_int_of)
  14.370 -done
  14.371 -
  14.372 -lemma not_zneg_nat_of_intify:
  14.373 -     "~ znegative(intify(z)) ==> $# (nat_of(z)) = intify(z)"
  14.374 -by (simp (no_asm_simp) add: nat_of_def not_zneg_raw_nat_of)
  14.375 -
  14.376 -lemma not_zneg_nat_of: "[| ~ znegative(z); z: int |] ==> $# (nat_of(z)) = z"
  14.377 -apply (simp (no_asm_simp) add: not_zneg_nat_of_intify)
  14.378 -done
  14.379 -
  14.380 -lemma zneg_nat_of [simp]: "znegative(intify(z)) ==> nat_of(z) = 0"
  14.381 -apply (subgoal_tac "intify(z) \<in> int")
  14.382 -apply (simp add: int_def) 
  14.383 -apply (auto simp add: znegative nat_of_def raw_nat_of 
  14.384 -            split add: nat_diff_split) 
  14.385 -done
  14.386 -
  14.387 -
  14.388 -subsection{*@{term zadd}: addition on int*}
  14.389 -
  14.390 -text{*Congruence Property for Addition*}
  14.391 -lemma zadd_congruent2: 
  14.392 -    "(%z1 z2. let <x1,y1>=z1; <x2,y2>=z2                  
  14.393 -                            in intrel``{<x1#+x2, y1#+y2>})
  14.394 -     respects2 intrel"
  14.395 -apply (simp add: congruent2_def)
  14.396 -(*Proof via congruent2_commuteI seems longer*)
  14.397 -apply safe
  14.398 -apply (simp (no_asm_simp) add: add_assoc Let_def)
  14.399 -(*The rest should be trivial, but rearranging terms is hard
  14.400 -  add_ac does not help rewriting with the assumptions.*)
  14.401 -apply (rule_tac m1 = x1a in add_left_commute [THEN ssubst])
  14.402 -apply (rule_tac m1 = x2a in add_left_commute [THEN ssubst])
  14.403 -apply (simp (no_asm_simp) add: add_assoc [symmetric])
  14.404 -done
  14.405 -
  14.406 -lemma raw_zadd_type: "[| z: int;  w: int |] ==> raw_zadd(z,w) : int"
  14.407 -apply (simp add: int_def raw_zadd_def)
  14.408 -apply (rule UN_equiv_class_type2 [OF equiv_intrel zadd_congruent2], assumption+)
  14.409 -apply (simp add: Let_def)
  14.410 -done
  14.411 -
  14.412 -lemma zadd_type [iff,TC]: "z $+ w : int"
  14.413 -by (simp add: zadd_def raw_zadd_type)
  14.414 -
  14.415 -lemma raw_zadd: 
  14.416 -  "[| x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat |]               
  14.417 -   ==> raw_zadd (intrel``{<x1,y1>}, intrel``{<x2,y2>}) =   
  14.418 -       intrel `` {<x1#+x2, y1#+y2>}"
  14.419 -apply (simp add: raw_zadd_def 
  14.420 -             UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2])
  14.421 -apply (simp add: Let_def)
  14.422 -done
  14.423 -
  14.424 -lemma zadd: 
  14.425 -  "[| x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat |]          
  14.426 -   ==> (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) =   
  14.427 -       intrel `` {<x1#+x2, y1#+y2>}"
  14.428 -by (simp add: zadd_def raw_zadd image_intrel_int)
  14.429 -
  14.430 -lemma raw_zadd_int0: "z : int ==> raw_zadd ($#0,z) = z"
  14.431 -by (auto simp add: int_def int_of_def raw_zadd)
  14.432 -
  14.433 -lemma zadd_int0_intify [simp]: "$#0 $+ z = intify(z)"
  14.434 -by (simp add: zadd_def raw_zadd_int0)
  14.435 -
  14.436 -lemma zadd_int0: "z: int ==> $#0 $+ z = z"
  14.437 -by simp
  14.438 -
  14.439 -lemma raw_zminus_zadd_distrib: 
  14.440 -     "[| z: int;  w: int |] ==> $- raw_zadd(z,w) = raw_zadd($- z, $- w)"
  14.441 -by (auto simp add: zminus raw_zadd int_def)
  14.442 -
  14.443 -lemma zminus_zadd_distrib [simp]: "$- (z $+ w) = $- z $+ $- w"
  14.444 -by (simp add: zadd_def raw_zminus_zadd_distrib)
  14.445 -
  14.446 -lemma raw_zadd_commute:
  14.447 -     "[| z: int;  w: int |] ==> raw_zadd(z,w) = raw_zadd(w,z)"
  14.448 -by (auto simp add: raw_zadd add_ac int_def)
  14.449 -
  14.450 -lemma zadd_commute: "z $+ w = w $+ z"
  14.451 -by (simp add: zadd_def raw_zadd_commute)
  14.452 -
  14.453 -lemma raw_zadd_assoc: 
  14.454 -    "[| z1: int;  z2: int;  z3: int |]    
  14.455 -     ==> raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))"
  14.456 -by (auto simp add: int_def raw_zadd add_assoc)
  14.457 -
  14.458 -lemma zadd_assoc: "(z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)"
  14.459 -by (simp add: zadd_def raw_zadd_type raw_zadd_assoc)
  14.460 -
  14.461 -(*For AC rewriting*)
  14.462 -lemma zadd_left_commute: "z1$+(z2$+z3) = z2$+(z1$+z3)"
  14.463 -apply (simp add: zadd_assoc [symmetric])
  14.464 -apply (simp add: zadd_commute)
  14.465 -done
  14.466 -
  14.467 -(*Integer addition is an AC operator*)
  14.468 -lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
  14.469 -
  14.470 -lemma int_of_add: "$# (m #+ n) = ($#m) $+ ($#n)"
  14.471 -by (simp add: int_of_def zadd)
  14.472 -
  14.473 -lemma int_succ_int_1: "$# succ(m) = $# 1 $+ ($# m)"
  14.474 -by (simp add: int_of_add [symmetric] natify_succ)
  14.475 -
  14.476 -lemma int_of_diff: 
  14.477 -     "[| m\<in>nat;  n le m |] ==> $# (m #- n) = ($#m) $- ($#n)"
  14.478 -apply (simp add: int_of_def zdiff_def)
  14.479 -apply (frule lt_nat_in_nat)
  14.480 -apply (simp_all add: zadd zminus add_diff_inverse2)
  14.481 -done
  14.482 -
  14.483 -lemma raw_zadd_zminus_inverse: "z : int ==> raw_zadd (z, $- z) = $#0"
  14.484 -by (auto simp add: int_def int_of_def zminus raw_zadd add_commute)
  14.485 -
  14.486 -lemma zadd_zminus_inverse [simp]: "z $+ ($- z) = $#0"
  14.487 -apply (simp add: zadd_def)
  14.488 -apply (subst zminus_intify [symmetric])
  14.489 -apply (rule intify_in_int [THEN raw_zadd_zminus_inverse])
  14.490 -done
  14.491 -
  14.492 -lemma zadd_zminus_inverse2 [simp]: "($- z) $+ z = $#0"
  14.493 -by (simp add: zadd_commute zadd_zminus_inverse)
  14.494 -
  14.495 -lemma zadd_int0_right_intify [simp]: "z $+ $#0 = intify(z)"
  14.496 -by (rule trans [OF zadd_commute zadd_int0_intify])
  14.497 -
  14.498 -lemma zadd_int0_right: "z:int ==> z $+ $#0 = z"
  14.499 -by simp
  14.500 -
  14.501 -
  14.502 -subsection{*@{term zmult}: Integer Multiplication*}
  14.503 -
  14.504 -text{*Congruence property for multiplication*}
  14.505 -lemma zmult_congruent2:
  14.506 -    "(%p1 p2. split(%x1 y1. split(%x2 y2.      
  14.507 -                    intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1))
  14.508 -     respects2 intrel"
  14.509 -apply (rule equiv_intrel [THEN congruent2_commuteI], auto)
  14.510 -(*Proof that zmult is congruent in one argument*)
  14.511 -apply (rename_tac x y)
  14.512 -apply (frule_tac t = "%u. x#*u" in sym [THEN subst_context])
  14.513 -apply (drule_tac t = "%u. y#*u" in subst_context)
  14.514 -apply (erule add_left_cancel)+
  14.515 -apply (simp_all add: add_mult_distrib_left)
  14.516 -done
  14.517 -
  14.518 -
  14.519 -lemma raw_zmult_type: "[| z: int;  w: int |] ==> raw_zmult(z,w) : int"
  14.520 -apply (simp add: int_def raw_zmult_def)
  14.521 -apply (rule UN_equiv_class_type2 [OF equiv_intrel zmult_congruent2], assumption+)
  14.522 -apply (simp add: Let_def)
  14.523 -done
  14.524 -
  14.525 -lemma zmult_type [iff,TC]: "z $* w : int"
  14.526 -by (simp add: zmult_def raw_zmult_type)
  14.527 -
  14.528 -lemma raw_zmult: 
  14.529 -     "[| x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat |]     
  14.530 -      ==> raw_zmult(intrel``{<x1,y1>}, intrel``{<x2,y2>}) =      
  14.531 -          intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}"
  14.532 -by (simp add: raw_zmult_def 
  14.533 -           UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2])
  14.534 -
  14.535 -lemma zmult: 
  14.536 -     "[| x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat |]     
  14.537 -      ==> (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) =      
  14.538 -          intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}"
  14.539 -by (simp add: zmult_def raw_zmult image_intrel_int)
  14.540 -
  14.541 -lemma raw_zmult_int0: "z : int ==> raw_zmult ($#0,z) = $#0"
  14.542 -by (auto simp add: int_def int_of_def raw_zmult)
  14.543 -
  14.544 -lemma zmult_int0 [simp]: "$#0 $* z = $#0"
  14.545 -by (simp add: zmult_def raw_zmult_int0)
  14.546 -
  14.547 -lemma raw_zmult_int1: "z : int ==> raw_zmult ($#1,z) = z"
  14.548 -by (auto simp add: int_def int_of_def raw_zmult)
  14.549 -
  14.550 -lemma zmult_int1_intify [simp]: "$#1 $* z = intify(z)"
  14.551 -by (simp add: zmult_def raw_zmult_int1)
  14.552 -
  14.553 -lemma zmult_int1: "z : int ==> $#1 $* z = z"
  14.554 -by simp
  14.555 -
  14.556 -lemma raw_zmult_commute:
  14.557 -     "[| z: int;  w: int |] ==> raw_zmult(z,w) = raw_zmult(w,z)"
  14.558 -by (auto simp add: int_def raw_zmult add_ac mult_ac)
  14.559 -
  14.560 -lemma zmult_commute: "z $* w = w $* z"
  14.561 -by (simp add: zmult_def raw_zmult_commute)
  14.562 -
  14.563 -lemma raw_zmult_zminus: 
  14.564 -     "[| z: int;  w: int |] ==> raw_zmult($- z, w) = $- raw_zmult(z, w)"
  14.565 -by (auto simp add: int_def zminus raw_zmult add_ac)
  14.566 -
  14.567 -lemma zmult_zminus [simp]: "($- z) $* w = $- (z $* w)"
  14.568 -apply (simp add: zmult_def raw_zmult_zminus)
  14.569 -apply (subst zminus_intify [symmetric], rule raw_zmult_zminus, auto)
  14.570 -done
  14.571 -
  14.572 -lemma zmult_zminus_right [simp]: "w $* ($- z) = $- (w $* z)"
  14.573 -by (simp add: zmult_commute [of w])
  14.574 -
  14.575 -lemma raw_zmult_assoc: 
  14.576 -    "[| z1: int;  z2: int;  z3: int |]    
  14.577 -     ==> raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))"
  14.578 -by (auto simp add: int_def raw_zmult add_mult_distrib_left add_ac mult_ac)
  14.579 -
  14.580 -lemma zmult_assoc: "(z1 $* z2) $* z3 = z1 $* (z2 $* z3)"
  14.581 -by (simp add: zmult_def raw_zmult_type raw_zmult_assoc)
  14.582 -
  14.583 -(*For AC rewriting*)
  14.584 -lemma zmult_left_commute: "z1$*(z2$*z3) = z2$*(z1$*z3)"
  14.585 -apply (simp add: zmult_assoc [symmetric])
  14.586 -apply (simp add: zmult_commute)
  14.587 -done
  14.588 -
  14.589 -(*Integer multiplication is an AC operator*)
  14.590 -lemmas zmult_ac = zmult_assoc zmult_commute zmult_left_commute
  14.591 -
  14.592 -lemma raw_zadd_zmult_distrib: 
  14.593 -    "[| z1: int;  z2: int;  w: int |]   
  14.594 -     ==> raw_zmult(raw_zadd(z1,z2), w) =  
  14.595 -         raw_zadd (raw_zmult(z1,w), raw_zmult(z2,w))"
  14.596 -by (auto simp add: int_def raw_zadd raw_zmult add_mult_distrib_left add_ac mult_ac)
  14.597 -
  14.598 -lemma zadd_zmult_distrib: "(z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)"
  14.599 -by (simp add: zmult_def zadd_def raw_zadd_type raw_zmult_type 
  14.600 -              raw_zadd_zmult_distrib)
  14.601 -
  14.602 -lemma zadd_zmult_distrib2: "w $* (z1 $+ z2) = (w $* z1) $+ (w $* z2)"
  14.603 -by (simp add: zmult_commute [of w] zadd_zmult_distrib)
  14.604 -
  14.605 -lemmas int_typechecks = 
  14.606 -  int_of_type zminus_type zmagnitude_type zadd_type zmult_type
  14.607 -
  14.608 -
  14.609 -(*** Subtraction laws ***)
  14.610 -
  14.611 -lemma zdiff_type [iff,TC]: "z $- w : int"
  14.612 -by (simp add: zdiff_def)
  14.613 -
  14.614 -lemma zminus_zdiff_eq [simp]: "$- (z $- y) = y $- z"
  14.615 -by (simp add: zdiff_def zadd_commute)
  14.616 -
  14.617 -lemma zdiff_zmult_distrib: "(z1 $- z2) $* w = (z1 $* w) $- (z2 $* w)"
  14.618 -apply (simp add: zdiff_def)
  14.619 -apply (subst zadd_zmult_distrib)
  14.620 -apply (simp add: zmult_zminus)
  14.621 -done
  14.622 -
  14.623 -lemma zdiff_zmult_distrib2: "w $* (z1 $- z2) = (w $* z1) $- (w $* z2)"
  14.624 -by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
  14.625 -
  14.626 -lemma zadd_zdiff_eq: "x $+ (y $- z) = (x $+ y) $- z"
  14.627 -by (simp add: zdiff_def zadd_ac)
  14.628 -
  14.629 -lemma zdiff_zadd_eq: "(x $- y) $+ z = (x $+ z) $- y"
  14.630 -by (simp add: zdiff_def zadd_ac)
  14.631 -
  14.632 -
  14.633 -subsection{*The "Less Than" Relation*}
  14.634 -
  14.635 -(*"Less than" is a linear ordering*)
  14.636 -lemma zless_linear_lemma: 
  14.637 -     "[| z: int; w: int |] ==> z$<w | z=w | w$<z"
  14.638 -apply (simp add: int_def zless_def znegative_def zdiff_def, auto)
  14.639 -apply (simp add: zadd zminus image_iff Bex_def)
  14.640 -apply (rule_tac i = "xb#+ya" and j = "xc #+ y" in Ord_linear_lt)
  14.641 -apply (force dest!: spec simp add: add_ac)+
  14.642 -done
  14.643 -
  14.644 -lemma zless_linear: "z$<w | intify(z)=intify(w) | w$<z"
  14.645 -apply (cut_tac z = " intify (z) " and w = " intify (w) " in zless_linear_lemma)
  14.646 -apply auto
  14.647 -done
  14.648 -
  14.649 -lemma zless_not_refl [iff]: "~ (z$<z)"
  14.650 -by (auto simp add: zless_def znegative_def int_of_def zdiff_def)
  14.651 -
  14.652 -lemma neq_iff_zless: "[| x: int; y: int |] ==> (x ~= y) <-> (x $< y | y $< x)"
  14.653 -by (cut_tac z = x and w = y in zless_linear, auto)
  14.654 -
  14.655 -lemma zless_imp_intify_neq: "w $< z ==> intify(w) ~= intify(z)"
  14.656 -apply auto
  14.657 -apply (subgoal_tac "~ (intify (w) $< intify (z))")
  14.658 -apply (erule_tac [2] ssubst)
  14.659 -apply (simp (no_asm_use))
  14.660 -apply auto
  14.661 -done
  14.662 -
  14.663 -(*This lemma allows direct proofs of other <-properties*)
  14.664 -lemma zless_imp_succ_zadd_lemma: 
  14.665 -    "[| w $< z; w: int; z: int |] ==> (\<exists>n\<in>nat. z = w $+ $#(succ(n)))"
  14.666 -apply (simp add: zless_def znegative_def zdiff_def int_def)
  14.667 -apply (auto dest!: less_imp_succ_add simp add: zadd zminus int_of_def)
  14.668 -apply (rule_tac x = k in bexI)
  14.669 -apply (erule add_left_cancel, auto)
  14.670 -done
  14.671 -
  14.672 -lemma zless_imp_succ_zadd:
  14.673 -     "w $< z ==> (\<exists>n\<in>nat. w $+ $#(succ(n)) = intify(z))"
  14.674 -apply (subgoal_tac "intify (w) $< intify (z) ")
  14.675 -apply (drule_tac w = "intify (w) " in zless_imp_succ_zadd_lemma)
  14.676 -apply auto
  14.677 -done
  14.678 -
  14.679 -lemma zless_succ_zadd_lemma: 
  14.680 -    "w : int ==> w $< w $+ $# succ(n)"
  14.681 -apply (simp add: zless_def znegative_def zdiff_def int_def)
  14.682 -apply (auto simp add: zadd zminus int_of_def image_iff)
  14.683 -apply (rule_tac x = 0 in exI, auto)
  14.684 -done
  14.685 -
  14.686 -lemma zless_succ_zadd: "w $< w $+ $# succ(n)"
  14.687 -by (cut_tac intify_in_int [THEN zless_succ_zadd_lemma], auto)
  14.688 -
  14.689 -lemma zless_iff_succ_zadd:
  14.690 -     "w $< z <-> (\<exists>n\<in>nat. w $+ $#(succ(n)) = intify(z))"
  14.691 -apply (rule iffI)
  14.692 -apply (erule zless_imp_succ_zadd, auto)
  14.693 -apply (rename_tac "n")
  14.694 -apply (cut_tac w = w and n = n in zless_succ_zadd, auto)
  14.695 -done
  14.696 -
  14.697 -lemma zless_int_of [simp]: "[| m\<in>nat; n\<in>nat |] ==> ($#m $< $#n) <-> (m<n)"
  14.698 -apply (simp add: less_iff_succ_add zless_iff_succ_zadd int_of_add [symmetric])
  14.699 -apply (blast intro: sym)
  14.700 -done
  14.701 -
  14.702 -lemma zless_trans_lemma: 
  14.703 -    "[| x $< y; y $< z; x: int; y : int; z: int |] ==> x $< z"
  14.704 -apply (simp add: zless_def znegative_def zdiff_def int_def)
  14.705 -apply (auto simp add: zadd zminus image_iff)
  14.706 -apply (rename_tac x1 x2 y1 y2)
  14.707 -apply (rule_tac x = "x1#+x2" in exI)
  14.708 -apply (rule_tac x = "y1#+y2" in exI)
  14.709 -apply (auto simp add: add_lt_mono)
  14.710 -apply (rule sym)
  14.711 -apply (erule add_left_cancel)+
  14.712 -apply auto
  14.713 -done
  14.714 -
  14.715 -lemma zless_trans: "[| x $< y; y $< z |] ==> x $< z"
  14.716 -apply (subgoal_tac "intify (x) $< intify (z) ")
  14.717 -apply (rule_tac [2] y = "intify (y) " in zless_trans_lemma)
  14.718 -apply auto
  14.719 -done
  14.720 -
  14.721 -lemma zless_not_sym: "z $< w ==> ~ (w $< z)"
  14.722 -by (blast dest: zless_trans)
  14.723 -
  14.724 -(* [| z $< w; ~ P ==> w $< z |] ==> P *)
  14.725 -lemmas zless_asym = zless_not_sym [THEN swap, standard]
  14.726 -
  14.727 -lemma zless_imp_zle: "z $< w ==> z $<= w"
  14.728 -by (simp add: zle_def)
  14.729 -
  14.730 -lemma zle_linear: "z $<= w | w $<= z"
  14.731 -apply (simp add: zle_def)
  14.732 -apply (cut_tac zless_linear, blast)
  14.733 -done
  14.734 -
  14.735 -
  14.736 -subsection{*Less Than or Equals*}
  14.737 -
  14.738 -lemma zle_refl: "z $<= z"
  14.739 -by (simp add: zle_def)
  14.740 -
  14.741 -lemma zle_eq_refl: "x=y ==> x $<= y"
  14.742 -by (simp add: zle_refl)
  14.743 -
  14.744 -lemma zle_anti_sym_intify: "[| x $<= y; y $<= x |] ==> intify(x) = intify(y)"
  14.745 -apply (simp add: zle_def, auto)
  14.746 -apply (blast dest: zless_trans)
  14.747 -done
  14.748 -
  14.749 -lemma zle_anti_sym: "[| x $<= y; y $<= x; x: int; y: int |] ==> x=y"
  14.750 -by (drule zle_anti_sym_intify, auto)
  14.751 -
  14.752 -lemma zle_trans_lemma:
  14.753 -     "[| x: int; y: int; z: int; x $<= y; y $<= z |] ==> x $<= z"
  14.754 -apply (simp add: zle_def, auto)
  14.755 -apply (blast intro: zless_trans)
  14.756 -done
  14.757 -
  14.758 -lemma zle_trans: "[| x $<= y; y $<= z |] ==> x $<= z"
  14.759 -apply (subgoal_tac "intify (x) $<= intify (z) ")
  14.760 -apply (rule_tac [2] y = "intify (y) " in zle_trans_lemma)
  14.761 -apply auto
  14.762 -done
  14.763 -
  14.764 -lemma zle_zless_trans: "[| i $<= j; j $< k |] ==> i $< k"
  14.765 -apply (auto simp add: zle_def)
  14.766 -apply (blast intro: zless_trans)
  14.767 -apply (simp add: zless_def zdiff_def zadd_def)
  14.768 -done
  14.769 -
  14.770 -lemma zless_zle_trans: "[| i $< j; j $<= k |] ==> i $< k"
  14.771 -apply (auto simp add: zle_def)
  14.772 -apply (blast intro: zless_trans)
  14.773 -apply (simp add: zless_def zdiff_def zminus_def)
  14.774 -done
  14.775 -
  14.776 -lemma not_zless_iff_zle: "~ (z $< w) <-> (w $<= z)"
  14.777 -apply (cut_tac z = z and w = w in zless_linear)
  14.778 -apply (auto dest: zless_trans simp add: zle_def)
  14.779 -apply (auto dest!: zless_imp_intify_neq)
  14.780 -done
  14.781 -
  14.782 -lemma not_zle_iff_zless: "~ (z $<= w) <-> (w $< z)"
  14.783 -by (simp add: not_zless_iff_zle [THEN iff_sym])
  14.784 -
  14.785 -
  14.786 -subsection{*More subtraction laws (for @{text zcompare_rls})*}
  14.787 -
  14.788 -lemma zdiff_zdiff_eq: "(x $- y) $- z = x $- (y $+ z)"
  14.789 -by (simp add: zdiff_def zadd_ac)
  14.790 -
  14.791 -lemma zdiff_zdiff_eq2: "x $- (y $- z) = (x $+ z) $- y"
  14.792 -by (simp add: zdiff_def zadd_ac)
  14.793 -
  14.794 -lemma zdiff_zless_iff: "(x$-y $< z) <-> (x $< z $+ y)"
  14.795 -by (simp add: zless_def zdiff_def zadd_ac)
  14.796 -
  14.797 -lemma zless_zdiff_iff: "(x $< z$-y) <-> (x $+ y $< z)"
  14.798 -by (simp add: zless_def zdiff_def zadd_ac)
  14.799 -
  14.800 -lemma zdiff_eq_iff: "[| x: int; z: int |] ==> (x$-y = z) <-> (x = z $+ y)"
  14.801 -by (auto simp add: zdiff_def zadd_assoc)
  14.802 -
  14.803 -lemma eq_zdiff_iff: "[| x: int; z: int |] ==> (x = z$-y) <-> (x $+ y = z)"
  14.804 -by (auto simp add: zdiff_def zadd_assoc)
  14.805 -
  14.806 -lemma zdiff_zle_iff_lemma:
  14.807 -     "[| x: int; z: int |] ==> (x$-y $<= z) <-> (x $<= z $+ y)"
  14.808 -by (auto simp add: zle_def zdiff_eq_iff zdiff_zless_iff)
  14.809 -
  14.810 -lemma zdiff_zle_iff: "(x$-y $<= z) <-> (x $<= z $+ y)"
  14.811 -by (cut_tac zdiff_zle_iff_lemma [OF intify_in_int intify_in_int], simp)
  14.812 -
  14.813 -lemma zle_zdiff_iff_lemma:
  14.814 -     "[| x: int; z: int |] ==>(x $<= z$-y) <-> (x $+ y $<= z)"
  14.815 -apply (auto simp add: zle_def zdiff_eq_iff zless_zdiff_iff)
  14.816 -apply (auto simp add: zdiff_def zadd_assoc)
  14.817 -done
  14.818 -
  14.819 -lemma zle_zdiff_iff: "(x $<= z$-y) <-> (x $+ y $<= z)"
  14.820 -by (cut_tac zle_zdiff_iff_lemma [ OF intify_in_int intify_in_int], simp)
  14.821 -
  14.822 -text{*This list of rewrites simplifies (in)equalities by bringing subtractions
  14.823 -  to the top and then moving negative terms to the other side.  
  14.824 -  Use with @{text zadd_ac}*}
  14.825 -lemmas zcompare_rls =
  14.826 -     zdiff_def [symmetric]
  14.827 -     zadd_zdiff_eq zdiff_zadd_eq zdiff_zdiff_eq zdiff_zdiff_eq2 
  14.828 -     zdiff_zless_iff zless_zdiff_iff zdiff_zle_iff zle_zdiff_iff 
  14.829 -     zdiff_eq_iff eq_zdiff_iff
  14.830 -
  14.831 -
  14.832 -subsection{*Monotonicity and Cancellation Results for Instantiation
  14.833 -     of the CancelNumerals Simprocs*}
  14.834 -
  14.835 -lemma zadd_left_cancel:
  14.836 -     "[| w: int; w': int |] ==> (z $+ w' = z $+ w) <-> (w' = w)"
  14.837 -apply safe
  14.838 -apply (drule_tac t = "%x. x $+ ($-z) " in subst_context)
  14.839 -apply (simp add: zadd_ac)
  14.840 -done
  14.841 -
  14.842 -lemma zadd_left_cancel_intify [simp]:
  14.843 -     "(z $+ w' = z $+ w) <-> intify(w') = intify(w)"
  14.844 -apply (rule iff_trans)
  14.845 -apply (rule_tac [2] zadd_left_cancel, auto)
  14.846 -done
  14.847 -
  14.848 -lemma zadd_right_cancel:
  14.849 -     "[| w: int; w': int |] ==> (w' $+ z = w $+ z) <-> (w' = w)"
  14.850 -apply safe
  14.851 -apply (drule_tac t = "%x. x $+ ($-z) " in subst_context)
  14.852 -apply (simp add: zadd_ac)
  14.853 -done
  14.854 -
  14.855 -lemma zadd_right_cancel_intify [simp]:
  14.856 -     "(w' $+ z = w $+ z) <-> intify(w') = intify(w)"
  14.857 -apply (rule iff_trans)
  14.858 -apply (rule_tac [2] zadd_right_cancel, auto)
  14.859 -done
  14.860 -
  14.861 -lemma zadd_right_cancel_zless [simp]: "(w' $+ z $< w $+ z) <-> (w' $< w)"
  14.862 -by (simp add: zdiff_zless_iff [THEN iff_sym] zdiff_def zadd_assoc)
  14.863 -
  14.864 -lemma zadd_left_cancel_zless [simp]: "(z $+ w' $< z $+ w) <-> (w' $< w)"
  14.865 -by (simp add: zadd_commute [of z] zadd_right_cancel_zless)
  14.866 -
  14.867 -lemma zadd_right_cancel_zle [simp]: "(w' $+ z $<= w $+ z) <-> w' $<= w"
  14.868 -by (simp add: zle_def)
  14.869 -
  14.870 -lemma zadd_left_cancel_zle [simp]: "(z $+ w' $<= z $+ w) <->  w' $<= w"
  14.871 -by (simp add: zadd_commute [of z]  zadd_right_cancel_zle)
  14.872 -
  14.873 -
  14.874 -(*"v $<= w ==> v$+z $<= w$+z"*)
  14.875 -lemmas zadd_zless_mono1 = zadd_right_cancel_zless [THEN iffD2, standard]
  14.876 -
  14.877 -(*"v $<= w ==> z$+v $<= z$+w"*)
  14.878 -lemmas zadd_zless_mono2 = zadd_left_cancel_zless [THEN iffD2, standard]
  14.879 -
  14.880 -(*"v $<= w ==> v$+z $<= w$+z"*)
  14.881 -lemmas zadd_zle_mono1 = zadd_right_cancel_zle [THEN iffD2, standard]
  14.882 -
  14.883 -(*"v $<= w ==> z$+v $<= z$+w"*)
  14.884 -lemmas zadd_zle_mono2 = zadd_left_cancel_zle [THEN iffD2, standard]
  14.885 -
  14.886 -lemma zadd_zle_mono: "[| w' $<= w; z' $<= z |] ==> w' $+ z' $<= w $+ z"
  14.887 -by (erule zadd_zle_mono1 [THEN zle_trans], simp)
  14.888 -
  14.889 -lemma zadd_zless_mono: "[| w' $< w; z' $<= z |] ==> w' $+ z' $< w $+ z"
  14.890 -by (erule zadd_zless_mono1 [THEN zless_zle_trans], simp)
  14.891 -
  14.892 -
  14.893 -subsection{*Comparison laws*}
  14.894 -
  14.895 -lemma zminus_zless_zminus [simp]: "($- x $< $- y) <-> (y $< x)"
  14.896 -by (simp add: zless_def zdiff_def zadd_ac)
  14.897 -
  14.898 -lemma zminus_zle_zminus [simp]: "($- x $<= $- y) <-> (y $<= x)"
  14.899 -by (simp add: not_zless_iff_zle [THEN iff_sym])
  14.900 -
  14.901 -subsubsection{*More inequality lemmas*}
  14.902 -
  14.903 -lemma equation_zminus: "[| x: int;  y: int |] ==> (x = $- y) <-> (y = $- x)"
  14.904 -by auto
  14.905 -
  14.906 -lemma zminus_equation: "[| x: int;  y: int |] ==> ($- x = y) <-> ($- y = x)"
  14.907 -by auto
  14.908 -
  14.909 -lemma equation_zminus_intify: "(intify(x) = $- y) <-> (intify(y) = $- x)"
  14.910 -apply (cut_tac x = "intify (x) " and y = "intify (y) " in equation_zminus)
  14.911 -apply auto
  14.912 -done
  14.913 -
  14.914 -lemma zminus_equation_intify: "($- x = intify(y)) <-> ($- y = intify(x))"
  14.915 -apply (cut_tac x = "intify (x) " and y = "intify (y) " in zminus_equation)
  14.916 -apply auto
  14.917 -done
  14.918 -
  14.919 -
  14.920 -subsubsection{*The next several equations are permutative: watch out!*}
  14.921 -
  14.922 -lemma zless_zminus: "(x $< $- y) <-> (y $< $- x)"
  14.923 -by (simp add: zless_def zdiff_def zadd_ac)
  14.924 -
  14.925 -lemma zminus_zless: "($- x $< y) <-> ($- y $< x)"
  14.926 -by (simp add: zless_def zdiff_def zadd_ac)
  14.927 -
  14.928 -lemma zle_zminus: "(x $<= $- y) <-> (y $<= $- x)"
  14.929 -by (simp add: not_zless_iff_zle [THEN iff_sym] zminus_zless)
  14.930 -
  14.931 -lemma zminus_zle: "($- x $<= y) <-> ($- y $<= x)"
  14.932 -by (simp add: not_zless_iff_zle [THEN iff_sym] zless_zminus)
  14.933 -
  14.934 -end
    15.1 --- a/src/ZF/IntDiv.thy	Mon Feb 11 15:19:17 2008 +0100
    15.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    15.3 @@ -1,1789 +0,0 @@
    15.4 -(*  Title:      ZF/IntDiv.thy
    15.5 -    ID:         $Id$
    15.6 -    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    15.7 -    Copyright   1999  University of Cambridge
    15.8 -
    15.9 -Here is the division algorithm in ML:
   15.10 -
   15.11 -    fun posDivAlg (a,b) =
   15.12 -      if a<b then (0,a)
   15.13 -      else let val (q,r) = posDivAlg(a, 2*b)
   15.14 -	       in  if 0<=r-b then (2*q+1, r-b) else (2*q, r)
   15.15 -	   end
   15.16 -
   15.17 -    fun negDivAlg (a,b) =
   15.18 -      if 0<=a+b then (~1,a+b)
   15.19 -      else let val (q,r) = negDivAlg(a, 2*b)
   15.20 -	       in  if 0<=r-b then (2*q+1, r-b) else (2*q, r)
   15.21 -	   end;
   15.22 -
   15.23 -    fun negateSnd (q,r:int) = (q,~r);
   15.24 -
   15.25 -    fun divAlg (a,b) = if 0<=a then 
   15.26 -			  if b>0 then posDivAlg (a,b) 
   15.27 -			   else if a=0 then (0,0)
   15.28 -				else negateSnd (negDivAlg (~a,~b))
   15.29 -		       else 
   15.30 -			  if 0<b then negDivAlg (a,b)
   15.31 -			  else        negateSnd (posDivAlg (~a,~b));
   15.32 -
   15.33 -*)
   15.34 -
   15.35 -header{*The Division Operators Div and Mod*}
   15.36 -
   15.37 -theory IntDiv imports IntArith OrderArith begin
   15.38 -
   15.39 -definition
   15.40 -  quorem :: "[i,i] => o"  where
   15.41 -    "quorem == %<a,b> <q,r>.
   15.42 -                      a = b$*q $+ r &
   15.43 -                      (#0$<b & #0$<=r & r$<b | ~(#0$<b) & b$<r & r $<= #0)"
   15.44 -
   15.45 -definition
   15.46 -  adjust :: "[i,i] => i"  where
   15.47 -    "adjust(b) == %<q,r>. if #0 $<= r$-b then <#2$*q $+ #1,r$-b>
   15.48 -                          else <#2$*q,r>"
   15.49 -
   15.50 -
   15.51 -(** the division algorithm **)
   15.52 -
   15.53 -definition
   15.54 -  posDivAlg :: "i => i"  where
   15.55 -(*for the case a>=0, b>0*)
   15.56 -(*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*)
   15.57 -    "posDivAlg(ab) ==
   15.58 -       wfrec(measure(int*int, %<a,b>. nat_of (a $- b $+ #1)),
   15.59 -	     ab,
   15.60 -	     %<a,b> f. if (a$<b | b$<=#0) then <#0,a>
   15.61 -                       else adjust(b, f ` <a,#2$*b>))"
   15.62 -
   15.63 -
   15.64 -(*for the case a<0, b>0*)
   15.65 -definition
   15.66 -  negDivAlg :: "i => i"  where
   15.67 -(*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*)
   15.68 -    "negDivAlg(ab) ==
   15.69 -       wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)),
   15.70 -	     ab,
   15.71 -	     %<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b>
   15.72 -                       else adjust(b, f ` <a,#2$*b>))"
   15.73 -
   15.74 -(*for the general case b\<noteq>0*)
   15.75 -
   15.76 -definition
   15.77 -  negateSnd :: "i => i"  where
   15.78 -    "negateSnd == %<q,r>. <q, $-r>"
   15.79 -
   15.80 -  (*The full division algorithm considers all possible signs for a, b
   15.81 -    including the special case a=0, b<0, because negDivAlg requires a<0*)
   15.82 -definition
   15.83 -  divAlg :: "i => i"  where
   15.84 -    "divAlg ==
   15.85 -       %<a,b>. if #0 $<= a then
   15.86 -                  if #0 $<= b then posDivAlg (<a,b>)
   15.87 -                  else if a=#0 then <#0,#0>
   15.88 -                       else negateSnd (negDivAlg (<$-a,$-b>))
   15.89 -               else 
   15.90 -                  if #0$<b then negDivAlg (<a,b>)
   15.91 -                  else         negateSnd (posDivAlg (<$-a,$-b>))"
   15.92 -
   15.93 -definition
   15.94 -  zdiv  :: "[i,i]=>i"                    (infixl "zdiv" 70)  where
   15.95 -    "a zdiv b == fst (divAlg (<intify(a), intify(b)>))"
   15.96 -
   15.97 -definition
   15.98 -  zmod  :: "[i,i]=>i"                    (infixl "zmod" 70)  where
   15.99 -    "a zmod b == snd (divAlg (<intify(a), intify(b)>))"
  15.100 -
  15.101 -
  15.102 -(** Some basic laws by Sidi Ehmety (need linear arithmetic!) **)
  15.103 -
  15.104 -lemma zspos_add_zspos_imp_zspos: "[| #0 $< x;  #0 $< y |] ==> #0 $< x $+ y"
  15.105 -apply (rule_tac y = "y" in zless_trans)
  15.106 -apply (rule_tac [2] zdiff_zless_iff [THEN iffD1])
  15.107 -apply auto
  15.108 -done
  15.109 -
  15.110 -lemma zpos_add_zpos_imp_zpos: "[| #0 $<= x;  #0 $<= y |] ==> #0 $<= x $+ y"
  15.111 -apply (rule_tac y = "y" in zle_trans)
  15.112 -apply (rule_tac [2] zdiff_zle_iff [THEN iffD1])
  15.113 -apply auto
  15.114 -done
  15.115 -
  15.116 -lemma zneg_add_zneg_imp_zneg: "[| x $< #0;  y $< #0 |] ==> x $+ y $< #0"
  15.117 -apply (rule_tac y = "y" in zless_trans)
  15.118 -apply (rule zless_zdiff_iff [THEN iffD1])
  15.119 -apply auto
  15.120 -done
  15.121 -
  15.122 -(* this theorem is used below *)
  15.123 -lemma zneg_or_0_add_zneg_or_0_imp_zneg_or_0:
  15.124 -     "[| x $<= #0;  y $<= #0 |] ==> x $+ y $<= #0"
  15.125 -apply (rule_tac y = "y" in zle_trans)
  15.126 -apply (rule zle_zdiff_iff [THEN iffD1])
  15.127 -apply auto
  15.128 -done
  15.129 -
  15.130 -lemma zero_lt_zmagnitude: "[| #0 $< k; k \<in> int |] ==> 0 < zmagnitude(k)"
  15.131 -apply (drule zero_zless_imp_znegative_zminus)
  15.132 -apply (drule_tac [2] zneg_int_of)
  15.133 -apply (auto simp add: zminus_equation [of k])
  15.134 -apply (subgoal_tac "0 < zmagnitude ($# succ (n))")
  15.135 - apply simp
  15.136 -apply (simp only: zmagnitude_int_of)
  15.137 -apply simp
  15.138 -done
  15.139 -
  15.140 -
  15.141 -(*** Inequality lemmas involving $#succ(m) ***)
  15.142 -
  15.143 -lemma zless_add_succ_iff:
  15.144 -     "(w $< z $+ $# succ(m)) <-> (w $< z $+ $#m | intify(w) = z $+ $#m)"
  15.145 -apply (auto simp add: zless_iff_succ_zadd zadd_assoc int_of_add [symmetric])
  15.146 -apply (rule_tac [3] x = "0" in bexI)
  15.147 -apply (cut_tac m = "m" in int_succ_int_1)
  15.148 -apply (cut_tac m = "n" in int_succ_int_1)
  15.149 -apply simp
  15.150 -apply (erule natE)
  15.151 -apply auto
  15.152 -apply (rule_tac x = "succ (n) " in bexI)
  15.153 -apply auto
  15.154 -done
  15.155 -
  15.156 -lemma zadd_succ_lemma:
  15.157 -     "z \<in> int ==> (w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)"
  15.158 -apply (simp only: not_zless_iff_zle [THEN iff_sym] zless_add_succ_iff)
  15.159 -apply (auto intro: zle_anti_sym elim: zless_asym
  15.160 -            simp add: zless_imp_zle not_zless_iff_zle)
  15.161 -done
  15.162 -
  15.163 -lemma zadd_succ_zle_iff: "(w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)"
  15.164 -apply (cut_tac z = "intify (z)" in zadd_succ_lemma)
  15.165 -apply auto
  15.166 -done
  15.167 -
  15.168 -(** Inequality reasoning **)
  15.169 -
  15.170 -lemma zless_add1_iff_zle: "(w $< z $+ #1) <-> (w$<=z)"
  15.171 -apply (subgoal_tac "#1 = $# 1")
  15.172 -apply (simp only: zless_add_succ_iff zle_def)
  15.173 -apply auto
  15.174 -done
  15.175 -
  15.176 -lemma add1_zle_iff: "(w $+ #1 $<= z) <-> (w $< z)"
  15.177 -apply (subgoal_tac "#1 = $# 1")
  15.178 -apply (simp only: zadd_succ_zle_iff)
  15.179 -apply auto
  15.180 -done
  15.181 -
  15.182 -lemma add1_left_zle_iff: "(#1 $+ w $<= z) <-> (w $< z)"
  15.183 -apply (subst zadd_commute)
  15.184 -apply (rule add1_zle_iff)
  15.185 -done
  15.186 -
  15.187 -
  15.188 -(*** Monotonicity of Multiplication ***)
  15.189 -
  15.190 -lemma zmult_mono_lemma: "k \<in> nat ==> i $<= j ==> i $* $#k $<= j $* $#k"
  15.191 -apply (induct_tac "k")
  15.192 - prefer 2 apply (subst int_succ_int_1)
  15.193 -apply (simp_all (no_asm_simp) add: zadd_zmult_distrib2 zadd_zle_mono)
  15.194 -done
  15.195 -
  15.196 -lemma zmult_zle_mono1: "[| i $<= j;  #0 $<= k |] ==> i$*k $<= j$*k"
  15.197 -apply (subgoal_tac "i $* intify (k) $<= j $* intify (k) ")
  15.198 -apply (simp (no_asm_use))
  15.199 -apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst])
  15.200 -apply (rule_tac [3] zmult_mono_lemma)
  15.201 -apply auto
  15.202 -apply (simp add: znegative_iff_zless_0 not_zless_iff_zle [THEN iff_sym])
  15.203 -done
  15.204 -
  15.205 -lemma zmult_zle_mono1_neg: "[| i $<= j;  k $<= #0 |] ==> j$*k $<= i$*k"
  15.206 -apply (rule zminus_zle_zminus [THEN iffD1])
  15.207 -apply (simp del: zmult_zminus_right
  15.208 -            add: zmult_zminus_right [symmetric] zmult_zle_mono1 zle_zminus)
  15.209 -done
  15.210 -
  15.211 -lemma zmult_zle_mono2: "[| i $<= j;  #0 $<= k |] ==> k$*i $<= k$*j"
  15.212 -apply (drule zmult_zle_mono1)
  15.213 -apply (simp_all add: zmult_commute)
  15.214 -done
  15.215 -
  15.216 -lemma zmult_zle_mono2_neg: "[| i $<= j;  k $<= #0 |] ==> k$*j $<= k$*i"
  15.217 -apply (drule zmult_zle_mono1_neg)
  15.218 -apply (simp_all add: zmult_commute)
  15.219 -done
  15.220 -
  15.221 -(* $<= monotonicity, BOTH arguments*)
  15.222 -lemma zmult_zle_mono:
  15.223 -     "[| i $<= j;  k $<= l;  #0 $<= j;  #0 $<= k |] ==> i$*k $<= j$*l"
  15.224 -apply (erule zmult_zle_mono1 [THEN zle_trans])
  15.225 -apply assumption
  15.226 -apply (erule zmult_zle_mono2)
  15.227 -apply assumption
  15.228 -done
  15.229 -
  15.230 -
  15.231 -(** strict, in 1st argument; proof is by induction on k>0 **)
  15.232 -
  15.233 -lemma zmult_zless_mono2_lemma [rule_format]:
  15.234 -     "[| i$<j; k \<in> nat |] ==> 0<k --> $#k $* i $< $#k $* j"
  15.235 -apply (induct_tac "k")
  15.236 - prefer 2
  15.237 - apply (subst int_succ_int_1)
  15.238 - apply (erule natE)
  15.239 -apply (simp_all add: zadd_zmult_distrib zadd_zless_mono zle_def)
  15.240 -apply (frule nat_0_le)
  15.241 -apply (subgoal_tac "i $+ (i $+ $# xa $* i) $< j $+ (j $+ $# xa $* j) ")
  15.242 -apply (simp (no_asm_use))
  15.243 -apply (rule zadd_zless_mono)
  15.244 -apply (simp_all (no_asm_simp) add: zle_def)
  15.245 -done
  15.246 -
  15.247 -lemma zmult_zless_mono2: "[| i$<j;  #0 $< k |] ==> k$*i $< k$*j"
  15.248 -apply (subgoal_tac "intify (k) $* i $< intify (k) $* j")
  15.249 -apply (simp (no_asm_use))
  15.250 -apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst])
  15.251 -apply (rule_tac [3] zmult_zless_mono2_lemma)
  15.252 -apply auto
  15.253 -apply (simp add: znegative_iff_zless_0)
  15.254 -apply (drule zless_trans, assumption)
  15.255 -apply (auto simp add: zero_lt_zmagnitude)
  15.256 -done
  15.257 -
  15.258 -lemma zmult_zless_mono1: "[| i$<j;  #0 $< k |] ==> i$*k $< j$*k"
  15.259 -apply (drule zmult_zless_mono2)
  15.260 -apply (simp_all add: zmult_commute)
  15.261 -done
  15.262 -
  15.263 -(* < monotonicity, BOTH arguments*)
  15.264 -lemma zmult_zless_mono:
  15.265 -     "[| i $< j;  k $< l;  #0 $< j;  #0 $< k |] ==> i$*k $< j$*l"
  15.266 -apply (erule zmult_zless_mono1 [THEN zless_trans])
  15.267 -apply assumption
  15.268 -apply (erule zmult_zless_mono2)
  15.269 -apply assumption
  15.270 -done
  15.271 -
  15.272 -lemma zmult_zless_mono1_neg: "[| i $< j;  k $< #0 |] ==> j$*k $< i$*k"
  15.273 -apply (rule zminus_zless_zminus [THEN iffD1])
  15.274 -apply (simp del: zmult_zminus_right 
  15.275 -            add: zmult_zminus_right [symmetric] zmult_zless_mono1 zless_zminus)
  15.276 -done
  15.277 -
  15.278 -lemma zmult_zless_mono2_neg: "[| i $< j;  k $< #0 |] ==> k$*j $< k$*i"
  15.279 -apply (rule zminus_zless_zminus [THEN iffD1])
  15.280 -apply (simp del: zmult_zminus 
  15.281 -            add: zmult_zminus [symmetric] zmult_zless_mono2 zless_zminus)
  15.282 -done
  15.283 -
  15.284 -
  15.285 -(** Products of zeroes **)
  15.286 -
  15.287 -lemma zmult_eq_lemma:
  15.288 -     "[| m \<in> int; n \<in> int |] ==> (m = #0 | n = #0) <-> (m$*n = #0)"
  15.289 -apply (case_tac "m $< #0")
  15.290 -apply (auto simp add: not_zless_iff_zle zle_def neq_iff_zless)
  15.291 -apply (force dest: zmult_zless_mono1_neg zmult_zless_mono1)+
  15.292 -done
  15.293 -
  15.294 -lemma zmult_eq_0_iff [iff]: "(m$*n = #0) <-> (intify(m) = #0 | intify(n) = #0)"
  15.295 -apply (simp add: zmult_eq_lemma)
  15.296 -done
  15.297 -
  15.298 -
  15.299 -(** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =,
  15.300 -    but not (yet?) for k*m < n*k. **)
  15.301 -
  15.302 -lemma zmult_zless_lemma:
  15.303 -     "[| k \<in> int; m \<in> int; n \<in> int |]  
  15.304 -      ==> (m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
  15.305 -apply (case_tac "k = #0")
  15.306 -apply (auto simp add: neq_iff_zless zmult_zless_mono1 zmult_zless_mono1_neg)
  15.307 -apply (auto simp add: not_zless_iff_zle 
  15.308 -                      not_zle_iff_zless [THEN iff_sym, of "m$*k"] 
  15.309 -                      not_zle_iff_zless [THEN iff_sym, of m])
  15.310 -apply (auto elim: notE
  15.311 -            simp add: zless_imp_zle zmult_zle_mono1 zmult_zle_mono1_neg)
  15.312 -done
  15.313 -
  15.314 -lemma zmult_zless_cancel2:
  15.315 -     "(m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
  15.316 -apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)" 
  15.317 -       in zmult_zless_lemma)
  15.318 -apply auto
  15.319 -done
  15.320 -
  15.321 -lemma zmult_zless_cancel1:
  15.322 -     "(k$*m $< k$*n) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
  15.323 -by (simp add: zmult_commute [of k] zmult_zless_cancel2)
  15.324 -
  15.325 -lemma zmult_zle_cancel2:
  15.326 -     "(m$*k $<= n$*k) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))"
  15.327 -by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel2)
  15.328 -
  15.329 -lemma zmult_zle_cancel1:
  15.330 -     "(k$*m $<= k$*n) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))"
  15.331 -by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel1)
  15.332 -
  15.333 -lemma int_eq_iff_zle: "[| m \<in> int; n \<in> int |] ==> m=n <-> (m $<= n & n $<= m)"
  15.334 -apply (blast intro: zle_refl zle_anti_sym)
  15.335 -done
  15.336 -
  15.337 -lemma zmult_cancel2_lemma:
  15.338 -     "[| k \<in> int; m \<in> int; n \<in> int |] ==> (m$*k = n$*k) <-> (k=#0 | m=n)"
  15.339 -apply (simp add: int_eq_iff_zle [of "m$*k"] int_eq_iff_zle [of m])
  15.340 -apply (auto simp add: zmult_zle_cancel2 neq_iff_zless)
  15.341 -done
  15.342 -
  15.343 -lemma zmult_cancel2 [simp]:
  15.344 -     "(m$*k = n$*k) <-> (intify(k) = #0 | intify(m) = intify(n))"
  15.345 -apply (rule iff_trans)
  15.346 -apply (rule_tac [2] zmult_cancel2_lemma)
  15.347 -apply auto
  15.348 -done
  15.349 -
  15.350 -lemma zmult_cancel1 [simp]:
  15.351 -     "(k$*m = k$*n) <-> (intify(k) = #0 | intify(m) = intify(n))"
  15.352 -by (simp add: zmult_commute [of k] zmult_cancel2)
  15.353 -
  15.354 -
  15.355 -subsection{* Uniqueness and monotonicity of quotients and remainders *}
  15.356 -
  15.357 -lemma unique_quotient_lemma:
  15.358 -     "[| b$*q' $+ r' $<= b$*q $+ r;  #0 $<= r';  #0 $< b;  r $< b |]  
  15.359 -      ==> q' $<= q"
  15.360 -apply (subgoal_tac "r' $+ b $* (q'$-q) $<= r")
  15.361 - prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_ac zcompare_rls)
  15.362 -apply (subgoal_tac "#0 $< b $* (#1 $+ q $- q') ")
  15.363 - prefer 2
  15.364 - apply (erule zle_zless_trans)
  15.365 - apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls)
  15.366 - apply (erule zle_zless_trans)
  15.367 - apply (simp add: ); 
  15.368 -apply (subgoal_tac "b $* q' $< b $* (#1 $+ q)")
  15.369 - prefer 2 
  15.370 - apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls)
  15.371 -apply (auto elim: zless_asym
  15.372 -        simp add: zmult_zless_cancel1 zless_add1_iff_zle zadd_ac zcompare_rls)
  15.373 -done
  15.374 -
  15.375 -lemma unique_quotient_lemma_neg:
  15.376 -     "[| b$*q' $+ r' $<= b$*q $+ r;  r $<= #0;  b $< #0;  b $< r' |]  
  15.377 -      ==> q $<= q'"
  15.378 -apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r" 
  15.379 -       in unique_quotient_lemma)
  15.380 -apply (auto simp del: zminus_zadd_distrib
  15.381 -            simp add: zminus_zadd_distrib [symmetric] zle_zminus zless_zminus)
  15.382 -done
  15.383 -
  15.384 -
  15.385 -lemma unique_quotient:
  15.386 -     "[| quorem (<a,b>, <q,r>);  quorem (<a,b>, <q',r'>);  b \<in> int; b ~= #0;  
  15.387 -         q \<in> int; q' \<in> int |] ==> q = q'"
  15.388 -apply (simp add: split_ifs quorem_def neq_iff_zless)
  15.389 -apply safe
  15.390 -apply simp_all
  15.391 -apply (blast intro: zle_anti_sym
  15.392 -             dest: zle_eq_refl [THEN unique_quotient_lemma] 
  15.393 -                   zle_eq_refl [THEN unique_quotient_lemma_neg] sym)+
  15.394 -done
  15.395 -
  15.396 -lemma unique_remainder:
  15.397 -     "[| quorem (<a,b>, <q,r>);  quorem (<a,b>, <q',r'>);  b \<in> int; b ~= #0;  
  15.398 -         q \<in> int; q' \<in> int;  
  15.399 -         r \<in> int; r' \<in> int |] ==> r = r'"
  15.400 -apply (subgoal_tac "q = q'")
  15.401 - prefer 2 apply (blast intro: unique_quotient)
  15.402 -apply (simp add: quorem_def)
  15.403 -done
  15.404 -
  15.405 -
  15.406 -subsection{*Correctness of posDivAlg, 
  15.407 -           the Division Algorithm for @{text "a\<ge>0"} and @{text "b>0"} *}
  15.408 -
  15.409 -lemma adjust_eq [simp]:
  15.410 -     "adjust(b, <q,r>) = (let diff = r$-b in  
  15.411 -                          if #0 $<= diff then <#2$*q $+ #1,diff>   
  15.412 -                                         else <#2$*q,r>)"
  15.413 -by (simp add: Let_def adjust_def)
  15.414 -
  15.415 -
  15.416 -lemma posDivAlg_termination:
  15.417 -     "[| #0 $< b; ~ a $< b |]    
  15.418 -      ==> nat_of(a $- #2 $\<times> b $+ #1) < nat_of(a $- b $+ #1)"
  15.419 -apply (simp (no_asm) add: zless_nat_conj)
  15.420 -apply (simp add: not_zless_iff_zle zless_add1_iff_zle zcompare_rls)
  15.421 -done
  15.422 -
  15.423 -lemmas posDivAlg_unfold = def_wfrec [OF posDivAlg_def wf_measure]
  15.424 -
  15.425 -lemma posDivAlg_eqn:
  15.426 -     "[| #0 $< b; a \<in> int; b \<in> int |] ==>  
  15.427 -      posDivAlg(<a,b>) =       
  15.428 -       (if a$<b then <#0,a> else adjust(b, posDivAlg (<a, #2$*b>)))"
  15.429 -apply (rule posDivAlg_unfold [THEN trans])
  15.430 -apply (simp add: vimage_iff not_zless_iff_zle [THEN iff_sym])
  15.431 -apply (blast intro: posDivAlg_termination)
  15.432 -done
  15.433 -
  15.434 -lemma posDivAlg_induct_lemma [rule_format]:
  15.435 -  assumes prem:
  15.436 -        "!!a b. [| a \<in> int; b \<in> int;  
  15.437 -                   ~ (a $< b | b $<= #0) --> P(<a, #2 $* b>) |] ==> P(<a,b>)"
  15.438 -  shows "<u,v> \<in> int*int --> P(<u,v>)"
  15.439 -apply (rule_tac a = "<u,v>" in wf_induct)
  15.440 -apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)" 
  15.441 -       in wf_measure)
  15.442 -apply clarify
  15.443 -apply (rule prem)
  15.444 -apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec)
  15.445 -apply auto
  15.446 -apply (simp add: not_zle_iff_zless posDivAlg_termination)
  15.447 -done
  15.448 -
  15.449 -
  15.450 -lemma posDivAlg_induct [consumes 2]:
  15.451 -  assumes u_int: "u \<in> int"
  15.452 -      and v_int: "v \<in> int"
  15.453 -      and ih: "!!a b. [| a \<in> int; b \<in> int;
  15.454 -                     ~ (a $< b | b $<= #0) --> P(a, #2 $* b) |] ==> P(a,b)"
  15.455 -  shows "P(u,v)"
  15.456 -apply (subgoal_tac "(%<x,y>. P (x,y)) (<u,v>)")
  15.457 -apply simp
  15.458 -apply (rule posDivAlg_induct_lemma)
  15.459 -apply (simp (no_asm_use))
  15.460 -apply (rule ih)
  15.461 -apply (auto simp add: u_int v_int)
  15.462 -done
  15.463 -
  15.464 -(*FIXME: use intify in integ_of so that we always have integ_of w \<in> int.
  15.465 -    then this rewrite can work for ALL constants!!*)
  15.466 -lemma intify_eq_0_iff_zle: "intify(m) = #0 <-> (m $<= #0 & #0 $<= m)"
  15.467 -apply (simp (no_asm) add: int_eq_iff_zle)
  15.468 -done
  15.469 -
  15.470 -
  15.471 -subsection{* Some convenient biconditionals for products of signs *}
  15.472 -
  15.473 -lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j"
  15.474 -apply (drule zmult_zless_mono1)
  15.475 -apply auto
  15.476 -done
  15.477 -
  15.478 -lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j"
  15.479 -apply (drule zmult_zless_mono1_neg)
  15.480 -apply auto
  15.481 -done
  15.482 -
  15.483 -lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0"
  15.484 -apply (drule zmult_zless_mono1_neg)
  15.485 -apply auto
  15.486 -done
  15.487 -
  15.488 -(** Inequality reasoning **)
  15.489 -
  15.490 -lemma int_0_less_lemma:
  15.491 -     "[| x \<in> int; y \<in> int |]  
  15.492 -      ==> (#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)"
  15.493 -apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg)
  15.494 -apply (rule ccontr) 
  15.495 -apply (rule_tac [2] ccontr) 
  15.496 -apply (auto simp add: zle_def not_zless_iff_zle)
  15.497 -apply (erule_tac P = "#0$< x$* y" in rev_mp)
  15.498 -apply (erule_tac [2] P = "#0$< x$* y" in rev_mp)
  15.499 -apply (drule zmult_pos_neg, assumption) 
  15.500 - prefer 2
  15.501 - apply (drule zmult_pos_neg, assumption) 
  15.502 -apply (auto dest: zless_not_sym simp add: zmult_commute)
  15.503 -done
  15.504 -
  15.505 -lemma int_0_less_mult_iff:
  15.506 -     "(#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)"
  15.507 -apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_less_lemma)
  15.508 -apply auto
  15.509 -done
  15.510 -
  15.511 -lemma int_0_le_lemma:
  15.512 -     "[| x \<in> int; y \<in> int |]  
  15.513 -      ==> (#0 $<= x $* y) <-> (#0 $<= x & #0 $<= y | x $<= #0 & y $<= #0)"
  15.514 -by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff)
  15.515 -
  15.516 -lemma int_0_le_mult_iff:
  15.517 -     "(#0 $<= x $* y) <-> ((#0 $<= x & #0 $<= y) | (x $<= #0 & y $<= #0))"
  15.518 -apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_le_lemma)
  15.519 -apply auto
  15.520 -done
  15.521 -
  15.522 -lemma zmult_less_0_iff:
  15.523 -     "(x $* y $< #0) <-> (#0 $< x & y $< #0 | x $< #0 & #0 $< y)"
  15.524 -apply (auto simp add: int_0_le_mult_iff not_zle_iff_zless [THEN iff_sym])
  15.525 -apply (auto dest: zless_not_sym simp add: not_zle_iff_zless)
  15.526 -done
  15.527 -
  15.528 -lemma zmult_le_0_iff:
  15.529 -     "(x $* y $<= #0) <-> (#0 $<= x & y $<= #0 | x $<= #0 & #0 $<= y)"
  15.530 -by (auto dest: zless_not_sym
  15.531 -         simp add: int_0_less_mult_iff not_zless_iff_zle [THEN iff_sym])
  15.532 -
  15.533 -
  15.534 -(*Typechecking for posDivAlg*)
  15.535 -lemma posDivAlg_type [rule_format]:
  15.536 -     "[| a \<in> int; b \<in> int |] ==> posDivAlg(<a,b>) \<in> int * int"
  15.537 -apply (rule_tac u = "a" and v = "b" in posDivAlg_induct)
  15.538 -apply assumption+
  15.539 -apply (case_tac "#0 $< ba")
  15.540 - apply (simp add: posDivAlg_eqn adjust_def integ_of_type 
  15.541 -             split add: split_if_asm)
  15.542 - apply clarify
  15.543 - apply (simp add: int_0_less_mult_iff not_zle_iff_zless)
  15.544 -apply (simp add: not_zless_iff_zle)
  15.545 -apply (subst posDivAlg_unfold)
  15.546 -apply simp
  15.547 -done
  15.548 -
  15.549 -(*Correctness of posDivAlg: it computes quotients correctly*)
  15.550 -lemma posDivAlg_correct [rule_format]:
  15.551 -     "[| a \<in> int; b \<in> int |]  
  15.552 -      ==> #0 $<= a --> #0 $< b --> quorem (<a,b>, posDivAlg(<a,b>))"
  15.553 -apply (rule_tac u = "a" and v = "b" in posDivAlg_induct)
  15.554 -apply auto
  15.555 -   apply (simp_all add: quorem_def)
  15.556 -   txt{*base case: a<b*}
  15.557 -   apply (simp add: posDivAlg_eqn)
  15.558 -  apply (simp add: not_zless_iff_zle [THEN iff_sym])
  15.559 - apply (simp add: int_0_less_mult_iff)
  15.560 -txt{*main argument*}
  15.561 -apply (subst posDivAlg_eqn)
  15.562 -apply (simp_all (no_asm_simp))
  15.563 -apply (erule splitE)
  15.564 -apply (rule posDivAlg_type)
  15.565 -apply (simp_all add: int_0_less_mult_iff)
  15.566 -apply (auto simp add: zadd_zmult_distrib2 Let_def)
  15.567 -txt{*now just linear arithmetic*}
  15.568 -apply (simp add: not_zle_iff_zless zdiff_zless_iff)
  15.569 -done
  15.570 -
  15.571 -
  15.572 -subsection{*Correctness of negDivAlg, the division algorithm for a<0 and b>0*}
  15.573 -
  15.574 -lemma negDivAlg_termination:
  15.575 -     "[| #0 $< b; a $+ b $< #0 |] 
  15.576 -      ==> nat_of($- a $- #2 $* b) < nat_of($- a $- b)"
  15.577 -apply (simp (no_asm) add: zless_nat_conj)
  15.578 -apply (simp add: zcompare_rls not_zle_iff_zless zless_zdiff_iff [THEN iff_sym]
  15.579 -                 zless_zminus)
  15.580 -done
  15.581 -
  15.582 -lemmas negDivAlg_unfold = def_wfrec [OF negDivAlg_def wf_measure]
  15.583 -
  15.584 -lemma negDivAlg_eqn:
  15.585 -     "[| #0 $< b; a : int; b : int |] ==>  
  15.586 -      negDivAlg(<a,b>) =       
  15.587 -       (if #0 $<= a$+b then <#-1,a$+b>  
  15.588 -                       else adjust(b, negDivAlg (<a, #2$*b>)))"
  15.589 -apply (rule negDivAlg_unfold [THEN trans])
  15.590 -apply (simp (no_asm_simp) add: vimage_iff not_zless_iff_zle [THEN iff_sym])
  15.591 -apply (blast intro: negDivAlg_termination)
  15.592 -done
  15.593 -
  15.594 -lemma negDivAlg_induct_lemma [rule_format]:
  15.595 -  assumes prem:
  15.596 -        "!!a b. [| a \<in> int; b \<in> int;  
  15.597 -                   ~ (#0 $<= a $+ b | b $<= #0) --> P(<a, #2 $* b>) |]  
  15.598 -                ==> P(<a,b>)"
  15.599 -  shows "<u,v> \<in> int*int --> P(<u,v>)"
  15.600 -apply (rule_tac a = "<u,v>" in wf_induct)
  15.601 -apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)" 
  15.602 -       in wf_measure)
  15.603 -apply clarify
  15.604 -apply (rule prem)
  15.605 -apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec)
  15.606 -apply auto
  15.607 -apply (simp add: not_zle_iff_zless negDivAlg_termination)
  15.608 -done
  15.609 -
  15.610 -lemma negDivAlg_induct [consumes 2]:
  15.611 -  assumes u_int: "u \<in> int"
  15.612 -      and v_int: "v \<in> int"
  15.613 -      and ih: "!!a b. [| a \<in> int; b \<in> int;  
  15.614 -                         ~ (#0 $<= a $+ b | b $<= #0) --> P(a, #2 $* b) |]  
  15.615 -                      ==> P(a,b)"
  15.616 -  shows "P(u,v)"
  15.617 -apply (subgoal_tac " (%<x,y>. P (x,y)) (<u,v>)")
  15.618 -apply simp
  15.619 -apply (rule negDivAlg_induct_lemma)
  15.620 -apply (simp (no_asm_use))
  15.621 -apply (rule ih)
  15.622 -apply (auto simp add: u_int v_int)
  15.623 -done
  15.624 -
  15.625 -
  15.626 -(*Typechecking for negDivAlg*)
  15.627 -lemma negDivAlg_type:
  15.628 -     "[| a \<in> int; b \<in> int |] ==> negDivAlg(<a,b>) \<in> int * int"
  15.629 -apply (rule_tac u = "a" and v = "b" in negDivAlg_induct)
  15.630 -apply assumption+
  15.631 -apply (case_tac "#0 $< ba")
  15.632 - apply (simp add: negDivAlg_eqn adjust_def integ_of_type 
  15.633 -             split add: split_if_asm)
  15.634 - apply clarify
  15.635 - apply (simp add: int_0_less_mult_iff not_zle_iff_zless)
  15.636 -apply (simp add: not_zless_iff_zle)
  15.637 -apply (subst negDivAlg_unfold)
  15.638 -apply simp
  15.639 -done
  15.640 -
  15.641 -
  15.642 -(*Correctness of negDivAlg: it computes quotients correctly
  15.643 -  It doesn't work if a=0 because the 0/b=0 rather than -1*)
  15.644 -lemma negDivAlg_correct [rule_format]:
  15.645 -     "[| a \<in> int; b \<in> int |]  
  15.646 -      ==> a $< #0 --> #0 $< b --> quorem (<a,b>, negDivAlg(<a,b>))"
  15.647 -apply (rule_tac u = "a" and v = "b" in negDivAlg_induct)
  15.648 -  apply auto
  15.649 -   apply (simp_all add: quorem_def)
  15.650 -   txt{*base case: @{term "0$<=a$+b"}*}
  15.651 -   apply (simp add: negDivAlg_eqn)
  15.652 -  apply (simp add: not_zless_iff_zle [THEN iff_sym])
  15.653 - apply (simp add: int_0_less_mult_iff)
  15.654 -txt{*main argument*}
  15.655 -apply (subst negDivAlg_eqn)
  15.656 -apply (simp_all (no_asm_simp))
  15.657 -apply (erule splitE)
  15.658 -apply (rule negDivAlg_type)
  15.659 -apply (simp_all add: int_0_less_mult_iff)
  15.660 -apply (auto simp add: zadd_zmult_distrib2 Let_def)
  15.661 -txt{*now just linear arithmetic*}
  15.662 -apply (simp add: not_zle_iff_zless zdiff_zless_iff)
  15.663 -done
  15.664 -
  15.665 -
  15.666 -subsection{* Existence shown by proving the division algorithm to be correct *}
  15.667 -
  15.668 -(*the case a=0*)
  15.669 -lemma quorem_0: "[|b \<noteq> #0;  b \<in> int|] ==> quorem (<#0,b>, <#0,#0>)"
  15.670 -by (force simp add: quorem_def neq_iff_zless)
  15.671 -
  15.672 -lemma posDivAlg_zero_divisor: "posDivAlg(<a,#0>) = <#0,a>"
  15.673 -apply (subst posDivAlg_unfold)
  15.674 -apply simp
  15.675 -done
  15.676 -
  15.677 -lemma posDivAlg_0 [simp]: "posDivAlg (<#0,b>) = <#0,#0>"
  15.678 -apply (subst posDivAlg_unfold)
  15.679 -apply (simp add: not_zle_iff_zless)
  15.680 -done
  15.681 -
  15.682 -
  15.683 -(*Needed below.  Actually it's an equivalence.*)
  15.684 -lemma linear_arith_lemma: "~ (#0 $<= #-1 $+ b) ==> (b $<= #0)"
  15.685 -apply (simp add: not_zle_iff_zless)
  15.686 -apply (drule zminus_zless_zminus [THEN iffD2])
  15.687 -apply (simp add: zadd_commute zless_add1_iff_zle zle_zminus)
  15.688 -done
  15.689 -
  15.690 -lemma negDivAlg_minus1 [simp]: "negDivAlg (<#-1,b>) = <#-1, b$-#1>"
  15.691 -apply (subst negDivAlg_unfold)
  15.692 -apply (simp add: linear_arith_lemma integ_of_type vimage_iff)
  15.693 -done
  15.694 -
  15.695 -lemma negateSnd_eq [simp]: "negateSnd (<q,r>) = <q, $-r>"
  15.696 -apply (unfold negateSnd_def)
  15.697 -apply auto
  15.698 -done
  15.699 -
  15.700 -lemma negateSnd_type: "qr \<in> int * int ==> negateSnd (qr) \<in> int * int"
  15.701 -apply (unfold negateSnd_def)
  15.702 -apply auto
  15.703 -done
  15.704 -
  15.705 -lemma quorem_neg:
  15.706 -     "[|quorem (<$-a,$-b>, qr);  a \<in> int;  b \<in> int;  qr \<in> int * int|]   
  15.707 -      ==> quorem (<a,b>, negateSnd(qr))"
  15.708 -apply clarify
  15.709 -apply (auto elim: zless_asym simp add: quorem_def zless_zminus)
  15.710 -txt{*linear arithmetic from here on*}
  15.711 -apply (simp_all add: zminus_equation [of a] zminus_zless)
  15.712 -apply (cut_tac [2] z = "b" and w = "#0" in zless_linear)
  15.713 -apply (cut_tac [1] z = "b" and w = "#0" in zless_linear)
  15.714 -apply auto
  15.715 -apply (blast dest: zle_zless_trans)+
  15.716 -done
  15.717 -
  15.718 -lemma divAlg_correct:
  15.719 -     "[|b \<noteq> #0;  a \<in> int;  b \<in> int|] ==> quorem (<a,b>, divAlg(<a,b>))"
  15.720 -apply (auto simp add: quorem_0 divAlg_def)
  15.721 -apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct
  15.722 -                    posDivAlg_type negDivAlg_type) 
  15.723 -apply (auto simp add: quorem_def neq_iff_zless)
  15.724 -txt{*linear arithmetic from here on*}
  15.725 -apply (auto simp add: zle_def)
  15.726 -done
  15.727 -
  15.728 -lemma divAlg_type: "[|a \<in> int;  b \<in> int|] ==> divAlg(<a,b>) \<in> int * int"
  15.729 -apply (auto simp add: divAlg_def)
  15.730 -apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type)
  15.731 -done
  15.732 -
  15.733 -
  15.734 -(** intify cancellation **)
  15.735 -
  15.736 -lemma zdiv_intify1 [simp]: "intify(x) zdiv y = x zdiv y"
  15.737 -apply (simp (no_asm) add: zdiv_def)
  15.738 -done
  15.739 -
  15.740 -lemma zdiv_intify2 [simp]: "x zdiv intify(y) = x zdiv y"
  15.741 -apply (simp (no_asm) add: zdiv_def)
  15.742 -done
  15.743 -
  15.744 -lemma zdiv_type [iff,TC]: "z zdiv w \<in> int"
  15.745 -apply (unfold zdiv_def)
  15.746 -apply (blast intro: fst_type divAlg_type)
  15.747 -done
  15.748 -
  15.749 -lemma zmod_intify1 [simp]: "intify(x) zmod y = x zmod y"
  15.750 -apply (simp (no_asm) add: zmod_def)
  15.751 -done
  15.752 -
  15.753 -lemma zmod_intify2 [simp]: "x zmod intify(y) = x zmod y"
  15.754 -apply (simp (no_asm) add: zmod_def)
  15.755 -done
  15.756 -
  15.757 -lemma zmod_type [iff,TC]: "z zmod w \<in> int"
  15.758 -apply (unfold zmod_def)
  15.759 -apply (rule snd_type)
  15.760 -apply (blast intro: divAlg_type)
  15.761 -done
  15.762 -
  15.763 -
  15.764 -(** Arbitrary definitions for division by zero.  Useful to simplify 
  15.765 -    certain equations **)
  15.766 -
  15.767 -lemma DIVISION_BY_ZERO_ZDIV: "a zdiv #0 = #0"
  15.768 -apply (simp (no_asm) add: zdiv_def divAlg_def posDivAlg_zero_divisor)
  15.769 -done  (*NOT for adding to default simpset*)
  15.770 -
  15.771 -lemma DIVISION_BY_ZERO_ZMOD: "a zmod #0 = intify(a)"
  15.772 -apply (simp (no_asm) add: zmod_def divAlg_def posDivAlg_zero_divisor)
  15.773 -done  (*NOT for adding to default simpset*)
  15.774 -
  15.775 -
  15.776 -
  15.777 -(** Basic laws about division and remainder **)
  15.778 -
  15.779 -lemma raw_zmod_zdiv_equality:
  15.780 -     "[| a \<in> int; b \<in> int |] ==> a = b $* (a zdiv b) $+ (a zmod b)"
  15.781 -apply (case_tac "b = #0")
  15.782 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
  15.783 -apply (cut_tac a = "a" and b = "b" in divAlg_correct)
  15.784 -apply (auto simp add: quorem_def zdiv_def zmod_def split_def)
  15.785 -done
  15.786 -
  15.787 -lemma zmod_zdiv_equality: "intify(a) = b $* (a zdiv b) $+ (a zmod b)"
  15.788 -apply (rule trans)
  15.789 -apply (rule_tac b = "intify (b)" in raw_zmod_zdiv_equality)
  15.790 -apply auto
  15.791 -done
  15.792 -
  15.793 -lemma pos_mod: "#0 $< b ==> #0 $<= a zmod b & a zmod b $< b"
  15.794 -apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct)
  15.795 -apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def)
  15.796 -apply (blast dest: zle_zless_trans)+
  15.797 -done
  15.798 -
  15.799 -lemmas pos_mod_sign = pos_mod [THEN conjunct1, standard]
  15.800 -and    pos_mod_bound = pos_mod [THEN conjunct2, standard]
  15.801 -
  15.802 -lemma neg_mod: "b $< #0 ==> a zmod b $<= #0 & b $< a zmod b"
  15.803 -apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct)
  15.804 -apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def)
  15.805 -apply (blast dest: zle_zless_trans)
  15.806 -apply (blast dest: zless_trans)+
  15.807 -done
  15.808 -
  15.809 -lemmas neg_mod_sign = neg_mod [THEN conjunct1, standard]
  15.810 -and    neg_mod_bound = neg_mod [THEN conjunct2, standard]
  15.811 -
  15.812 -
  15.813 -(** proving general properties of zdiv and zmod **)
  15.814 -
  15.815 -lemma quorem_div_mod:
  15.816 -     "[|b \<noteq> #0;  a \<in> int;  b \<in> int |]  
  15.817 -      ==> quorem (<a,b>, <a zdiv b, a zmod b>)"
  15.818 -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
  15.819 -apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound 
  15.820 -                      neg_mod_sign neg_mod_bound)
  15.821 -done
  15.822 -
  15.823 -(*Surely quorem(<a,b>,<q,r>) implies a \<in> int, but it doesn't matter*)
  15.824 -lemma quorem_div:
  15.825 -     "[| quorem(<a,b>,<q,r>);  b \<noteq> #0;  a \<in> int;  b \<in> int;  q \<in> int |]  
  15.826 -      ==> a zdiv b = q"
  15.827 -by (blast intro: quorem_div_mod [THEN unique_quotient])
  15.828 -
  15.829 -lemma quorem_mod:
  15.830 -     "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int |] 
  15.831 -      ==> a zmod b = r"
  15.832 -by (blast intro: quorem_div_mod [THEN unique_remainder])
  15.833 -
  15.834 -lemma zdiv_pos_pos_trivial_raw:
  15.835 -     "[| a \<in> int;  b \<in> int;  #0 $<= a;  a $< b |] ==> a zdiv b = #0"
  15.836 -apply (rule quorem_div)
  15.837 -apply (auto simp add: quorem_def)
  15.838 -(*linear arithmetic*)
  15.839 -apply (blast dest: zle_zless_trans)+
  15.840 -done
  15.841 -
  15.842 -lemma zdiv_pos_pos_trivial: "[| #0 $<= a;  a $< b |] ==> a zdiv b = #0"
  15.843 -apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  15.844 -       in zdiv_pos_pos_trivial_raw)
  15.845 -apply auto
  15.846 -done
  15.847 -
  15.848 -lemma zdiv_neg_neg_trivial_raw:
  15.849 -     "[| a \<in> int;  b \<in> int;  a $<= #0;  b $< a |] ==> a zdiv b = #0"
  15.850 -apply (rule_tac r = "a" in quorem_div)
  15.851 -apply (auto simp add: quorem_def)
  15.852 -(*linear arithmetic*)
  15.853 -apply (blast dest: zle_zless_trans zless_trans)+
  15.854 -done
  15.855 -
  15.856 -lemma zdiv_neg_neg_trivial: "[| a $<= #0;  b $< a |] ==> a zdiv b = #0"
  15.857 -apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  15.858 -       in zdiv_neg_neg_trivial_raw)
  15.859 -apply auto
  15.860 -done
  15.861 -
  15.862 -lemma zadd_le_0_lemma: "[| a$+b $<= #0;  #0 $< a;  #0 $< b |] ==> False"
  15.863 -apply (drule_tac z' = "#0" and z = "b" in zadd_zless_mono)
  15.864 -apply (auto simp add: zle_def)
  15.865 -apply (blast dest: zless_trans)
  15.866 -done
  15.867 -
  15.868 -lemma zdiv_pos_neg_trivial_raw:
  15.869 -     "[| a \<in> int;  b \<in> int;  #0 $< a;  a$+b $<= #0 |] ==> a zdiv b = #-1"
  15.870 -apply (rule_tac r = "a $+ b" in quorem_div)
  15.871 -apply (auto simp add: quorem_def)
  15.872 -(*linear arithmetic*)
  15.873 -apply (blast dest: zadd_le_0_lemma zle_zless_trans)+
  15.874 -done
  15.875 -
  15.876 -lemma zdiv_pos_neg_trivial: "[| #0 $< a;  a$+b $<= #0 |] ==> a zdiv b = #-1"
  15.877 -apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  15.878 -       in zdiv_pos_neg_trivial_raw)
  15.879 -apply auto
  15.880 -done
  15.881 -
  15.882 -(*There is no zdiv_neg_pos_trivial because  #0 zdiv b = #0 would supersede it*)
  15.883 -
  15.884 -
  15.885 -lemma zmod_pos_pos_trivial_raw:
  15.886 -     "[| a \<in> int;  b \<in> int;  #0 $<= a;  a $< b |] ==> a zmod b = a"
  15.887 -apply (rule_tac q = "#0" in quorem_mod)
  15.888 -apply (auto simp add: quorem_def)
  15.889 -(*linear arithmetic*)
  15.890 -apply (blast dest: zle_zless_trans)+
  15.891 -done
  15.892 -
  15.893 -lemma zmod_pos_pos_trivial: "[| #0 $<= a;  a $< b |] ==> a zmod b = intify(a)"
  15.894 -apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  15.895 -       in zmod_pos_pos_trivial_raw)
  15.896 -apply auto
  15.897 -done
  15.898 -
  15.899 -lemma zmod_neg_neg_trivial_raw:
  15.900 -     "[| a \<in> int;  b \<in> int;  a $<= #0;  b $< a |] ==> a zmod b = a"
  15.901 -apply (rule_tac q = "#0" in quorem_mod)
  15.902 -apply (auto simp add: quorem_def)
  15.903 -(*linear arithmetic*)
  15.904 -apply (blast dest: zle_zless_trans zless_trans)+
  15.905 -done
  15.906 -
  15.907 -lemma zmod_neg_neg_trivial: "[| a $<= #0;  b $< a |] ==> a zmod b = intify(a)"
  15.908 -apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  15.909 -       in zmod_neg_neg_trivial_raw)
  15.910 -apply auto
  15.911 -done
  15.912 -
  15.913 -lemma zmod_pos_neg_trivial_raw:
  15.914 -     "[| a \<in> int;  b \<in> int;  #0 $< a;  a$+b $<= #0 |] ==> a zmod b = a$+b"
  15.915 -apply (rule_tac q = "#-1" in quorem_mod)
  15.916 -apply (auto simp add: quorem_def)
  15.917 -(*linear arithmetic*)
  15.918 -apply (blast dest: zadd_le_0_lemma zle_zless_trans)+
  15.919 -done
  15.920 -
  15.921 -lemma zmod_pos_neg_trivial: "[| #0 $< a;  a$+b $<= #0 |] ==> a zmod b = a$+b"
  15.922 -apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  15.923 -       in zmod_pos_neg_trivial_raw)
  15.924 -apply auto
  15.925 -done
  15.926 -
  15.927 -(*There is no zmod_neg_pos_trivial...*)
  15.928 -
  15.929 -
  15.930 -(*Simpler laws such as -a zdiv b = -(a zdiv b) FAIL*)
  15.931 -
  15.932 -lemma zdiv_zminus_zminus_raw:
  15.933 -     "[|a \<in> int;  b \<in> int|] ==> ($-a) zdiv ($-b) = a zdiv b"
  15.934 -apply (case_tac "b = #0")
  15.935 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
  15.936 -apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_div])
  15.937 -apply auto
  15.938 -done
  15.939 -
  15.940 -lemma zdiv_zminus_zminus [simp]: "($-a) zdiv ($-b) = a zdiv b"
  15.941 -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zminus_zminus_raw)
  15.942 -apply auto
  15.943 -done
  15.944 -
  15.945 -(*Simpler laws such as -a zmod b = -(a zmod b) FAIL*)
  15.946 -lemma zmod_zminus_zminus_raw:
  15.947 -     "[|a \<in> int;  b \<in> int|] ==> ($-a) zmod ($-b) = $- (a zmod b)"
  15.948 -apply (case_tac "b = #0")
  15.949 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
  15.950 -apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod])
  15.951 -apply auto
  15.952 -done
  15.953 -
  15.954 -lemma zmod_zminus_zminus [simp]: "($-a) zmod ($-b) = $- (a zmod b)"
  15.955 -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zminus_zminus_raw)
  15.956 -apply auto
  15.957 -done
  15.958 -
  15.959 -
  15.960 -subsection{* division of a number by itself *}
  15.961 -
  15.962 -lemma self_quotient_aux1: "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $<= q"
  15.963 -apply (subgoal_tac "#0 $< a$*q")
  15.964 -apply (cut_tac w = "#0" and z = "q" in add1_zle_iff)
  15.965 -apply (simp add: int_0_less_mult_iff)
  15.966 -apply (blast dest: zless_trans)
  15.967 -(*linear arithmetic...*)
  15.968 -apply (drule_tac t = "%x. x $- r" in subst_context)
  15.969 -apply (drule sym)
  15.970 -apply (simp add: zcompare_rls)
  15.971 -done
  15.972 -
  15.973 -lemma self_quotient_aux2: "[| #0 $< a; a = r $+ a$*q; #0 $<= r |] ==> q $<= #1"
  15.974 -apply (subgoal_tac "#0 $<= a$* (#1$-q)")
  15.975 - apply (simp add: int_0_le_mult_iff zcompare_rls)
  15.976 - apply (blast dest: zle_zless_trans)
  15.977 -apply (simp add: zdiff_zmult_distrib2)
  15.978 -apply (drule_tac t = "%x. x $- a $* q" in subst_context)
  15.979 -apply (simp add: zcompare_rls)
  15.980 -done
  15.981 -
  15.982 -lemma self_quotient:
  15.983 -     "[| quorem(<a,a>,<q,r>);  a \<in> int;  q \<in> int;  a \<noteq> #0|] ==> q = #1"
  15.984 -apply (simp add: split_ifs quorem_def neq_iff_zless)
  15.985 -apply (rule zle_anti_sym)
  15.986 -apply safe
  15.987 -apply auto
  15.988 -prefer 4 apply (blast dest: zless_trans)
  15.989 -apply (blast dest: zless_trans)
  15.990 -apply (rule_tac [3] a = "$-a" and r = "$-r" in self_quotient_aux1)
  15.991 -apply (rule_tac a = "$-a" and r = "$-r" in self_quotient_aux2)
  15.992 -apply (rule_tac [6] zminus_equation [THEN iffD1])
  15.993 -apply (rule_tac [2] zminus_equation [THEN iffD1])
  15.994 -apply (force intro: self_quotient_aux1 self_quotient_aux2
  15.995 -  simp add: zadd_commute zmult_zminus)+
  15.996 -done
  15.997 -
  15.998 -lemma self_remainder:
  15.999 -     "[|quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0|] ==> r = #0"
 15.1000 -apply (frule self_quotient)
 15.1001 -apply (auto simp add: quorem_def)
 15.1002 -done
 15.1003 -
 15.1004 -lemma zdiv_self_raw: "[|a \<noteq> #0; a \<in> int|] ==> a zdiv a = #1"
 15.1005 -apply (blast intro: quorem_div_mod [THEN self_quotient])
 15.1006 -done
 15.1007 -
 15.1008 -lemma zdiv_self [simp]: "intify(a) \<noteq> #0 ==> a zdiv a = #1"
 15.1009 -apply (drule zdiv_self_raw)
 15.1010 -apply auto
 15.1011 -done
 15.1012 -
 15.1013 -(*Here we have 0 zmod 0 = 0, also assumed by Knuth (who puts m zmod 0 = 0) *)
 15.1014 -lemma zmod_self_raw: "a \<in> int ==> a zmod a = #0"
 15.1015 -apply (case_tac "a = #0")
 15.1016 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1017 -apply (blast intro: quorem_div_mod [THEN self_remainder])
 15.1018 -done
 15.1019 -
 15.1020 -lemma zmod_self [simp]: "a zmod a = #0"
 15.1021 -apply (cut_tac a = "intify (a)" in zmod_self_raw)
 15.1022 -apply auto
 15.1023 -done
 15.1024 -
 15.1025 -
 15.1026 -subsection{* Computation of division and remainder *}
 15.1027 -
 15.1028 -lemma zdiv_zero [simp]: "#0 zdiv b = #0"
 15.1029 -apply (simp (no_asm) add: zdiv_def divAlg_def)
 15.1030 -done
 15.1031 -
 15.1032 -lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1"
 15.1033 -apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
 15.1034 -done
 15.1035 -
 15.1036 -lemma zmod_zero [simp]: "#0 zmod b = #0"
 15.1037 -apply (simp (no_asm) add: zmod_def divAlg_def)
 15.1038 -done
 15.1039 -
 15.1040 -lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1"
 15.1041 -apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
 15.1042 -done
 15.1043 -
 15.1044 -lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1"
 15.1045 -apply (simp (no_asm_simp) add: zmod_def divAlg_def)
 15.1046 -done
 15.1047 -
 15.1048 -(** a positive, b positive **)
 15.1049 -
 15.1050 -lemma zdiv_pos_pos: "[| #0 $< a;  #0 $<= b |]  
 15.1051 -      ==> a zdiv b = fst (posDivAlg(<intify(a), intify(b)>))"
 15.1052 -apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
 15.1053 -apply (auto simp add: zle_def)
 15.1054 -done
 15.1055 -
 15.1056 -lemma zmod_pos_pos:
 15.1057 -     "[| #0 $< a;  #0 $<= b |]  
 15.1058 -      ==> a zmod b = snd (posDivAlg(<intify(a), intify(b)>))"
 15.1059 -apply (simp (no_asm_simp) add: zmod_def divAlg_def)
 15.1060 -apply (auto simp add: zle_def)
 15.1061 -done
 15.1062 -
 15.1063 -(** a negative, b positive **)
 15.1064 -
 15.1065 -lemma zdiv_neg_pos:
 15.1066 -     "[| a $< #0;  #0 $< b |]  
 15.1067 -      ==> a zdiv b = fst (negDivAlg(<intify(a), intify(b)>))"
 15.1068 -apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
 15.1069 -apply (blast dest: zle_zless_trans)
 15.1070 -done
 15.1071 -
 15.1072 -lemma zmod_neg_pos:
 15.1073 -     "[| a $< #0;  #0 $< b |]  
 15.1074 -      ==> a zmod b = snd (negDivAlg(<intify(a), intify(b)>))"
 15.1075 -apply (simp (no_asm_simp) add: zmod_def divAlg_def)
 15.1076 -apply (blast dest: zle_zless_trans)
 15.1077 -done
 15.1078 -
 15.1079 -(** a positive, b negative **)
 15.1080 -
 15.1081 -lemma zdiv_pos_neg:
 15.1082 -     "[| #0 $< a;  b $< #0 |]  
 15.1083 -      ==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))"
 15.1084 -apply (simp (no_asm_simp) add: zdiv_def divAlg_def intify_eq_0_iff_zle)
 15.1085 -apply auto
 15.1086 -apply (blast dest: zle_zless_trans)+
 15.1087 -apply (blast dest: zless_trans)
 15.1088 -apply (blast intro: zless_imp_zle)
 15.1089 -done
 15.1090 -
 15.1091 -lemma zmod_pos_neg:
 15.1092 -     "[| #0 $< a;  b $< #0 |]  
 15.1093 -      ==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))"
 15.1094 -apply (simp (no_asm_simp) add: zmod_def divAlg_def intify_eq_0_iff_zle)
 15.1095 -apply auto
 15.1096 -apply (blast dest: zle_zless_trans)+
 15.1097 -apply (blast dest: zless_trans)
 15.1098 -apply (blast intro: zless_imp_zle)
 15.1099 -done
 15.1100 -
 15.1101 -(** a negative, b negative **)
 15.1102 -
 15.1103 -lemma zdiv_neg_neg:
 15.1104 -     "[| a $< #0;  b $<= #0 |]  
 15.1105 -      ==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))"
 15.1106 -apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
 15.1107 -apply auto
 15.1108 -apply (blast dest!: zle_zless_trans)+
 15.1109 -done
 15.1110 -
 15.1111 -lemma zmod_neg_neg:
 15.1112 -     "[| a $< #0;  b $<= #0 |]  
 15.1113 -      ==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))"
 15.1114 -apply (simp (no_asm_simp) add: zmod_def divAlg_def)
 15.1115 -apply auto
 15.1116 -apply (blast dest!: zle_zless_trans)+
 15.1117 -done
 15.1118 -
 15.1119 -declare zdiv_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp]
 15.1120 -declare zdiv_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp]
 15.1121 -declare zdiv_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp]
 15.1122 -declare zdiv_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp]
 15.1123 -declare zmod_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp]
 15.1124 -declare zmod_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp]
 15.1125 -declare zmod_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp]
 15.1126 -declare zmod_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp]
 15.1127 -declare posDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp]
 15.1128 -declare negDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp]
 15.1129 -
 15.1130 -
 15.1131 -(** Special-case simplification **)
 15.1132 -
 15.1133 -lemma zmod_1 [simp]: "a zmod #1 = #0"
 15.1134 -apply (cut_tac a = "a" and b = "#1" in pos_mod_sign)
 15.1135 -apply (cut_tac [2] a = "a" and b = "#1" in pos_mod_bound)
 15.1136 -apply auto
 15.1137 -(*arithmetic*)
 15.1138 -apply (drule add1_zle_iff [THEN iffD2])
 15.1139 -apply (rule zle_anti_sym)
 15.1140 -apply auto
 15.1141 -done
 15.1142 -
 15.1143 -lemma zdiv_1 [simp]: "a zdiv #1 = intify(a)"
 15.1144 -apply (cut_tac a = "a" and b = "#1" in zmod_zdiv_equality)
 15.1145 -apply auto
 15.1146 -done
 15.1147 -
 15.1148 -lemma zmod_minus1_right [simp]: "a zmod #-1 = #0"
 15.1149 -apply (cut_tac a = "a" and b = "#-1" in neg_mod_sign)
 15.1150 -apply (cut_tac [2] a = "a" and b = "#-1" in neg_mod_bound)
 15.1151 -apply auto
 15.1152 -(*arithmetic*)
 15.1153 -apply (drule add1_zle_iff [THEN iffD2])
 15.1154 -apply (rule zle_anti_sym)
 15.1155 -apply auto
 15.1156 -done
 15.1157 -
 15.1158 -lemma zdiv_minus1_right_raw: "a \<in> int ==> a zdiv #-1 = $-a"
 15.1159 -apply (cut_tac a = "a" and b = "#-1" in zmod_zdiv_equality)
 15.1160 -apply auto
 15.1161 -apply (rule equation_zminus [THEN iffD2])
 15.1162 -apply auto
 15.1163 -done
 15.1164 -
 15.1165 -lemma zdiv_minus1_right: "a zdiv #-1 = $-a"
 15.1166 -apply (cut_tac a = "intify (a)" in zdiv_minus1_right_raw)
 15.1167 -apply auto
 15.1168 -done
 15.1169 -declare zdiv_minus1_right [simp]
 15.1170 -
 15.1171 -
 15.1172 -subsection{* Monotonicity in the first argument (divisor) *}
 15.1173 -
 15.1174 -lemma zdiv_mono1: "[| a $<= a';  #0 $< b |] ==> a zdiv b $<= a' zdiv b"
 15.1175 -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
 15.1176 -apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality)
 15.1177 -apply (rule unique_quotient_lemma)
 15.1178 -apply (erule subst)
 15.1179 -apply (erule subst)
 15.1180 -apply (simp_all (no_asm_simp) add: pos_mod_sign pos_mod_bound)
 15.1181 -done
 15.1182 -
 15.1183 -lemma zdiv_mono1_neg: "[| a $<= a';  b $< #0 |] ==> a' zdiv b $<= a zdiv b"
 15.1184 -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
 15.1185 -apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality)
 15.1186 -apply (rule unique_quotient_lemma_neg)
 15.1187 -apply (erule subst)
 15.1188 -apply (erule subst)
 15.1189 -apply (simp_all (no_asm_simp) add: neg_mod_sign neg_mod_bound)
 15.1190 -done
 15.1191 -
 15.1192 -
 15.1193 -subsection{* Monotonicity in the second argument (dividend) *}
 15.1194 -
 15.1195 -lemma q_pos_lemma:
 15.1196 -     "[| #0 $<= b'$*q' $+ r'; r' $< b';  #0 $< b' |] ==> #0 $<= q'"
 15.1197 -apply (subgoal_tac "#0 $< b'$* (q' $+ #1)")
 15.1198 - apply (simp add: int_0_less_mult_iff)
 15.1199 - apply (blast dest: zless_trans intro: zless_add1_iff_zle [THEN iffD1])
 15.1200 -apply (simp add: zadd_zmult_distrib2)
 15.1201 -apply (erule zle_zless_trans)
 15.1202 -apply (erule zadd_zless_mono2)
 15.1203 -done
 15.1204 -
 15.1205 -lemma zdiv_mono2_lemma:
 15.1206 -     "[| b$*q $+ r = b'$*q' $+ r';  #0 $<= b'$*q' $+ r';   
 15.1207 -         r' $< b';  #0 $<= r;  #0 $< b';  b' $<= b |]   
 15.1208 -      ==> q $<= q'"
 15.1209 -apply (frule q_pos_lemma, assumption+) 
 15.1210 -apply (subgoal_tac "b$*q $< b$* (q' $+ #1)")
 15.1211 - apply (simp add: zmult_zless_cancel1)
 15.1212 - apply (force dest: zless_add1_iff_zle [THEN iffD1] zless_trans zless_zle_trans)
 15.1213 -apply (subgoal_tac "b$*q = r' $- r $+ b'$*q'")
 15.1214 - prefer 2 apply (simp add: zcompare_rls)
 15.1215 -apply (simp (no_asm_simp) add: zadd_zmult_distrib2)
 15.1216 -apply (subst zadd_commute [of "b $\<times> q'"], rule zadd_zless_mono)
 15.1217 - prefer 2 apply (blast intro: zmult_zle_mono1)
 15.1218 -apply (subgoal_tac "r' $+ #0 $< b $+ r")
 15.1219 - apply (simp add: zcompare_rls)
 15.1220 -apply (rule zadd_zless_mono)
 15.1221 - apply auto
 15.1222 -apply (blast dest: zless_zle_trans)
 15.1223 -done
 15.1224 -
 15.1225 -
 15.1226 -lemma zdiv_mono2_raw:
 15.1227 -     "[| #0 $<= a;  #0 $< b';  b' $<= b;  a \<in> int |]   
 15.1228 -      ==> a zdiv b $<= a zdiv b'"
 15.1229 -apply (subgoal_tac "#0 $< b")
 15.1230 - prefer 2 apply (blast dest: zless_zle_trans)
 15.1231 -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
 15.1232 -apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality)
 15.1233 -apply (rule zdiv_mono2_lemma)
 15.1234 -apply (erule subst)
 15.1235 -apply (erule subst)
 15.1236 -apply (simp_all add: pos_mod_sign pos_mod_bound)
 15.1237 -done
 15.1238 -
 15.1239 -lemma zdiv_mono2:
 15.1240 -     "[| #0 $<= a;  #0 $< b';  b' $<= b |]   
 15.1241 -      ==> a zdiv b $<= a zdiv b'"
 15.1242 -apply (cut_tac a = "intify (a)" in zdiv_mono2_raw)
 15.1243 -apply auto
 15.1244 -done
 15.1245 -
 15.1246 -lemma q_neg_lemma:
 15.1247 -     "[| b'$*q' $+ r' $< #0;  #0 $<= r';  #0 $< b' |] ==> q' $< #0"
 15.1248 -apply (subgoal_tac "b'$*q' $< #0")
 15.1249 - prefer 2 apply (force intro: zle_zless_trans)
 15.1250 -apply (simp add: zmult_less_0_iff)
 15.1251 -apply (blast dest: zless_trans)
 15.1252 -done
 15.1253 -
 15.1254 -
 15.1255 -
 15.1256 -lemma zdiv_mono2_neg_lemma:
 15.1257 -     "[| b$*q $+ r = b'$*q' $+ r';  b'$*q' $+ r' $< #0;   
 15.1258 -         r $< b;  #0 $<= r';  #0 $< b';  b' $<= b |]   
 15.1259 -      ==> q' $<= q"
 15.1260 -apply (subgoal_tac "#0 $< b")
 15.1261 - prefer 2 apply (blast dest: zless_zle_trans)
 15.1262 -apply (frule q_neg_lemma, assumption+) 
 15.1263 -apply (subgoal_tac "b$*q' $< b$* (q $+ #1)")
 15.1264 - apply (simp add: zmult_zless_cancel1)
 15.1265 - apply (blast dest: zless_trans zless_add1_iff_zle [THEN iffD1])
 15.1266 -apply (simp (no_asm_simp) add: zadd_zmult_distrib2)
 15.1267 -apply (subgoal_tac "b$*q' $<= b'$*q'")
 15.1268 - prefer 2
 15.1269 - apply (simp add: zmult_zle_cancel2)
 15.1270 - apply (blast dest: zless_trans)
 15.1271 -apply (subgoal_tac "b'$*q' $+ r $< b $+ (b$*q $+ r)")
 15.1272 - prefer 2
 15.1273 - apply (erule ssubst)
 15.1274 - apply simp
 15.1275 - apply (drule_tac w' = "r" and z' = "#0" in zadd_zless_mono)
 15.1276 -  apply (assumption)
 15.1277 - apply simp
 15.1278 -apply (simp (no_asm_use) add: zadd_commute)
 15.1279 -apply (rule zle_zless_trans)
 15.1280 - prefer 2 apply (assumption)
 15.1281 -apply (simp (no_asm_simp) add: zmult_zle_cancel2)
 15.1282 -apply (blast dest: zless_trans)
 15.1283 -done
 15.1284 -
 15.1285 -lemma zdiv_mono2_neg_raw:
 15.1286 -     "[| a $< #0;  #0 $< b';  b' $<= b;  a \<in> int |]   
 15.1287 -      ==> a zdiv b' $<= a zdiv b"
 15.1288 -apply (subgoal_tac "#0 $< b")
 15.1289 - prefer 2 apply (blast dest: zless_zle_trans)
 15.1290 -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
 15.1291 -apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality)
 15.1292 -apply (rule zdiv_mono2_neg_lemma)
 15.1293 -apply (erule subst)
 15.1294 -apply (erule subst)
 15.1295 -apply (simp_all add: pos_mod_sign pos_mod_bound)
 15.1296 -done
 15.1297 -
 15.1298 -lemma zdiv_mono2_neg: "[| a $< #0;  #0 $< b';  b' $<= b |]   
 15.1299 -      ==> a zdiv b' $<= a zdiv b"
 15.1300 -apply (cut_tac a = "intify (a)" in zdiv_mono2_neg_raw)
 15.1301 -apply auto
 15.1302 -done
 15.1303 -
 15.1304 -
 15.1305 -
 15.1306 -subsection{* More algebraic laws for zdiv and zmod *}
 15.1307 -
 15.1308 -(** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **)
 15.1309 -
 15.1310 -lemma zmult1_lemma:
 15.1311 -     "[| quorem(<b,c>, <q,r>);  c \<in> int;  c \<noteq> #0 |]  
 15.1312 -      ==> quorem (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)"
 15.1313 -apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2
 15.1314 -                      pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound)
 15.1315 -apply (auto intro: raw_zmod_zdiv_equality) 
 15.1316 -done
 15.1317 -
 15.1318 -lemma zdiv_zmult1_eq_raw:
 15.1319 -     "[|b \<in> int;  c \<in> int|]  
 15.1320 -      ==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c"
 15.1321 -apply (case_tac "c = #0")
 15.1322 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1323 -apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_div])
 15.1324 -apply auto
 15.1325 -done
 15.1326 -
 15.1327 -lemma zdiv_zmult1_eq: "(a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c"
 15.1328 -apply (cut_tac b = "intify (b)" and c = "intify (c)" in zdiv_zmult1_eq_raw)
 15.1329 -apply auto
 15.1330 -done
 15.1331 -
 15.1332 -lemma zmod_zmult1_eq_raw:
 15.1333 -     "[|b \<in> int;  c \<in> int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c"
 15.1334 -apply (case_tac "c = #0")
 15.1335 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1336 -apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod])
 15.1337 -apply auto
 15.1338 -done
 15.1339 -
 15.1340 -lemma zmod_zmult1_eq: "(a$*b) zmod c = a$*(b zmod c) zmod c"
 15.1341 -apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult1_eq_raw)
 15.1342 -apply auto
 15.1343 -done
 15.1344 -
 15.1345 -lemma zmod_zmult1_eq': "(a$*b) zmod c = ((a zmod c) $* b) zmod c"
 15.1346 -apply (rule trans)
 15.1347 -apply (rule_tac b = " (b $* a) zmod c" in trans)
 15.1348 -apply (rule_tac [2] zmod_zmult1_eq)
 15.1349 -apply (simp_all (no_asm) add: zmult_commute)
 15.1350 -done
 15.1351 -
 15.1352 -lemma zmod_zmult_distrib: "(a$*b) zmod c = ((a zmod c) $* (b zmod c)) zmod c"
 15.1353 -apply (rule zmod_zmult1_eq' [THEN trans])
 15.1354 -apply (rule zmod_zmult1_eq)
 15.1355 -done
 15.1356 -
 15.1357 -lemma zdiv_zmult_self1 [simp]: "intify(b) \<noteq> #0 ==> (a$*b) zdiv b = intify(a)"
 15.1358 -apply (simp (no_asm_simp) add: zdiv_zmult1_eq)
 15.1359 -done
 15.1360 -
 15.1361 -lemma zdiv_zmult_self2 [simp]: "intify(b) \<noteq> #0 ==> (b$*a) zdiv b = intify(a)"
 15.1362 -apply (subst zmult_commute , erule zdiv_zmult_self1)
 15.1363 -done
 15.1364 -
 15.1365 -lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0"
 15.1366 -apply (simp (no_asm) add: zmod_zmult1_eq)
 15.1367 -done
 15.1368 -
 15.1369 -lemma zmod_zmult_self2 [simp]: "(b$*a) zmod b = #0"
 15.1370 -apply (simp (no_asm) add: zmult_commute zmod_zmult1_eq)
 15.1371 -done
 15.1372 -
 15.1373 -
 15.1374 -(** proving (a$+b) zdiv c = 
 15.1375 -            a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **)
 15.1376 -
 15.1377 -lemma zadd1_lemma:
 15.1378 -     "[| quorem(<a,c>, <aq,ar>);  quorem(<b,c>, <bq,br>);   
 15.1379 -         c \<in> int;  c \<noteq> #0 |]  
 15.1380 -      ==> quorem (<a$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)"
 15.1381 -apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2
 15.1382 -                      pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound)
 15.1383 -apply (auto intro: raw_zmod_zdiv_equality)
 15.1384 -done
 15.1385 -
 15.1386 -(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
 15.1387 -lemma zdiv_zadd1_eq_raw:
 15.1388 -     "[|a \<in> int; b \<in> int; c \<in> int|] ==>  
 15.1389 -      (a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)"
 15.1390 -apply (case_tac "c = #0")
 15.1391 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1392 -apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod,
 15.1393 -                                 THEN quorem_div])
 15.1394 -done
 15.1395 -
 15.1396 -lemma zdiv_zadd1_eq:
 15.1397 -     "(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)"
 15.1398 -apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" 
 15.1399 -       in zdiv_zadd1_eq_raw)
 15.1400 -apply auto
 15.1401 -done
 15.1402 -
 15.1403 -lemma zmod_zadd1_eq_raw:
 15.1404 -     "[|a \<in> int; b \<in> int; c \<in> int|]   
 15.1405 -      ==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c"
 15.1406 -apply (case_tac "c = #0")
 15.1407 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1408 -apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, 
 15.1409 -                                 THEN quorem_mod])
 15.1410 -done
 15.1411 -
 15.1412 -lemma zmod_zadd1_eq: "(a$+b) zmod c = (a zmod c $+ b zmod c) zmod c"
 15.1413 -apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" 
 15.1414 -       in zmod_zadd1_eq_raw)
 15.1415 -apply auto
 15.1416 -done
 15.1417 -
 15.1418 -lemma zmod_div_trivial_raw:
 15.1419 -     "[|a \<in> int; b \<in> int|] ==> (a zmod b) zdiv b = #0"
 15.1420 -apply (case_tac "b = #0")
 15.1421 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1422 -apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound
 15.1423 -         zdiv_pos_pos_trivial neg_mod_sign neg_mod_bound zdiv_neg_neg_trivial)
 15.1424 -done
 15.1425 -
 15.1426 -lemma zmod_div_trivial [simp]: "(a zmod b) zdiv b = #0"
 15.1427 -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_div_trivial_raw)
 15.1428 -apply auto
 15.1429 -done
 15.1430 -
 15.1431 -lemma zmod_mod_trivial_raw:
 15.1432 -     "[|a \<in> int; b \<in> int|] ==> (a zmod b) zmod b = a zmod b"
 15.1433 -apply (case_tac "b = #0")
 15.1434 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1435 -apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound 
 15.1436 -       zmod_pos_pos_trivial neg_mod_sign neg_mod_bound zmod_neg_neg_trivial)
 15.1437 -done
 15.1438 -
 15.1439 -lemma zmod_mod_trivial [simp]: "(a zmod b) zmod b = a zmod b"
 15.1440 -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_mod_trivial_raw)
 15.1441 -apply auto
 15.1442 -done
 15.1443 -
 15.1444 -lemma zmod_zadd_left_eq: "(a$+b) zmod c = ((a zmod c) $+ b) zmod c"
 15.1445 -apply (rule trans [symmetric])
 15.1446 -apply (rule zmod_zadd1_eq)
 15.1447 -apply (simp (no_asm))
 15.1448 -apply (rule zmod_zadd1_eq [symmetric])
 15.1449 -done
 15.1450 -
 15.1451 -lemma zmod_zadd_right_eq: "(a$+b) zmod c = (a $+ (b zmod c)) zmod c"
 15.1452 -apply (rule trans [symmetric])
 15.1453 -apply (rule zmod_zadd1_eq)
 15.1454 -apply (simp (no_asm))
 15.1455 -apply (rule zmod_zadd1_eq [symmetric])
 15.1456 -done
 15.1457 -
 15.1458 -
 15.1459 -lemma zdiv_zadd_self1 [simp]:
 15.1460 -     "intify(a) \<noteq> #0 ==> (a$+b) zdiv a = b zdiv a $+ #1"
 15.1461 -by (simp (no_asm_simp) add: zdiv_zadd1_eq)
 15.1462 -
 15.1463 -lemma zdiv_zadd_self2 [simp]:
 15.1464 -     "intify(a) \<noteq> #0 ==> (b$+a) zdiv a = b zdiv a $+ #1"
 15.1465 -by (simp (no_asm_simp) add: zdiv_zadd1_eq)
 15.1466 -
 15.1467 -lemma zmod_zadd_self1 [simp]: "(a$+b) zmod a = b zmod a"
 15.1468 -apply (case_tac "a = #0")
 15.1469 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1470 -apply (simp (no_asm_simp) add: zmod_zadd1_eq)
 15.1471 -done
 15.1472 -
 15.1473 -lemma zmod_zadd_self2 [simp]: "(b$+a) zmod a = b zmod a"
 15.1474 -apply (case_tac "a = #0")
 15.1475 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1476 -apply (simp (no_asm_simp) add: zmod_zadd1_eq)
 15.1477 -done
 15.1478 -
 15.1479 -
 15.1480 -subsection{* proving  a zdiv (b*c) = (a zdiv b) zdiv c *}
 15.1481 -
 15.1482 -(*The condition c>0 seems necessary.  Consider that 7 zdiv ~6 = ~2 but
 15.1483 -  7 zdiv 2 zdiv ~3 = 3 zdiv ~3 = ~1.  The subcase (a zdiv b) zmod c = 0 seems
 15.1484 -  to cause particular problems.*)
 15.1485 -
 15.1486 -(** first, four lemmas to bound the remainder for the cases b<0 and b>0 **)
 15.1487 -
 15.1488 -lemma zdiv_zmult2_aux1:
 15.1489 -     "[| #0 $< c;  b $< r;  r $<= #0 |] ==> b$*c $< b$*(q zmod c) $+ r"
 15.1490 -apply (subgoal_tac "b $* (c $- q zmod c) $< r $* #1")
 15.1491 -apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls)
 15.1492 -apply (rule zle_zless_trans)
 15.1493 -apply (erule_tac [2] zmult_zless_mono1)
 15.1494 -apply (rule zmult_zle_mono2_neg)
 15.1495 -apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound)
 15.1496 -apply (blast intro: zless_imp_zle dest: zless_zle_trans)
 15.1497 -done
 15.1498 -
 15.1499 -lemma zdiv_zmult2_aux2:
 15.1500 -     "[| #0 $< c;   b $< r;  r $<= #0 |] ==> b $* (q zmod c) $+ r $<= #0"
 15.1501 -apply (subgoal_tac "b $* (q zmod c) $<= #0")
 15.1502 - prefer 2
 15.1503 - apply (simp add: zmult_le_0_iff pos_mod_sign) 
 15.1504 - apply (blast intro: zless_imp_zle dest: zless_zle_trans)
 15.1505 -(*arithmetic*)
 15.1506 -apply (drule zadd_zle_mono)
 15.1507 -apply assumption
 15.1508 -apply (simp add: zadd_commute)
 15.1509 -done
 15.1510 -
 15.1511 -lemma zdiv_zmult2_aux3:
 15.1512 -     "[| #0 $< c;  #0 $<= r;  r $< b |] ==> #0 $<= b $* (q zmod c) $+ r"
 15.1513 -apply (subgoal_tac "#0 $<= b $* (q zmod c)")
 15.1514 - prefer 2
 15.1515 - apply (simp add: int_0_le_mult_iff pos_mod_sign) 
 15.1516 - apply (blast intro: zless_imp_zle dest: zle_zless_trans)
 15.1517 -(*arithmetic*)
 15.1518 -apply (drule zadd_zle_mono)
 15.1519 -apply assumption
 15.1520 -apply (simp add: zadd_commute)
 15.1521 -done
 15.1522 -
 15.1523 -lemma zdiv_zmult2_aux4:
 15.1524 -     "[| #0 $< c; #0 $<= r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c"
 15.1525 -apply (subgoal_tac "r $* #1 $< b $* (c $- q zmod c)")
 15.1526 -apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls)
 15.1527 -apply (rule zless_zle_trans)
 15.1528 -apply (erule zmult_zless_mono1)
 15.1529 -apply (rule_tac [2] zmult_zle_mono2)
 15.1530 -apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound)
 15.1531 -apply (blast intro: zless_imp_zle dest: zle_zless_trans)
 15.1532 -done
 15.1533 -
 15.1534 -lemma zdiv_zmult2_lemma:
 15.1535 -     "[| quorem (<a,b>, <q,r>);  a \<in> int;  b \<in> int;  b \<noteq> #0;  #0 $< c |]  
 15.1536 -      ==> quorem (<a,b$*c>, <q zdiv c, b$*(q zmod c) $+ r>)"
 15.1537 -apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def
 15.1538 -               neq_iff_zless int_0_less_mult_iff 
 15.1539 -               zadd_zmult_distrib2 [symmetric] zdiv_zmult2_aux1 zdiv_zmult2_aux2
 15.1540 -               zdiv_zmult2_aux3 zdiv_zmult2_aux4)
 15.1541 -apply (blast dest: zless_trans)+
 15.1542 -done
 15.1543 -
 15.1544 -lemma zdiv_zmult2_eq_raw:
 15.1545 -     "[|#0 $< c;  a \<in> int;  b \<in> int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c"
 15.1546 -apply (case_tac "b = #0")
 15.1547 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1548 -apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_div])
 15.1549 -apply (auto simp add: intify_eq_0_iff_zle)
 15.1550 -apply (blast dest: zle_zless_trans)
 15.1551 -done
 15.1552 -
 15.1553 -lemma zdiv_zmult2_eq: "#0 $< c ==> a zdiv (b$*c) = (a zdiv b) zdiv c"
 15.1554 -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zmult2_eq_raw)
 15.1555 -apply auto
 15.1556 -done
 15.1557 -
 15.1558 -lemma zmod_zmult2_eq_raw:
 15.1559 -     "[|#0 $< c;  a \<in> int;  b \<in> int|]  
 15.1560 -      ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b"
 15.1561 -apply (case_tac "b = #0")
 15.1562 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1563 -apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_mod])
 15.1564 -apply (auto simp add: intify_eq_0_iff_zle)
 15.1565 -apply (blast dest: zle_zless_trans)
 15.1566 -done
 15.1567 -
 15.1568 -lemma zmod_zmult2_eq:
 15.1569 -     "#0 $< c ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b"
 15.1570 -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zmult2_eq_raw)
 15.1571 -apply auto
 15.1572 -done
 15.1573 -
 15.1574 -subsection{* Cancellation of common factors in "zdiv" *}
 15.1575 -
 15.1576 -lemma zdiv_zmult_zmult1_aux1:
 15.1577 -     "[| #0 $< b;  intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b"
 15.1578 -apply (subst zdiv_zmult2_eq)
 15.1579 -apply auto
 15.1580 -done
 15.1581 -
 15.1582 -lemma zdiv_zmult_zmult1_aux2:
 15.1583 -     "[| b $< #0;  intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b"
 15.1584 -apply (subgoal_tac " (c $* ($-a)) zdiv (c $* ($-b)) = ($-a) zdiv ($-b)")
 15.1585 -apply (rule_tac [2] zdiv_zmult_zmult1_aux1)
 15.1586 -apply auto
 15.1587 -done
 15.1588 -
 15.1589 -lemma zdiv_zmult_zmult1_raw:
 15.1590 -     "[|intify(c) \<noteq> #0; b \<in> int|] ==> (c$*a) zdiv (c$*b) = a zdiv b"
 15.1591 -apply (case_tac "b = #0")
 15.1592 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1593 -apply (auto simp add: neq_iff_zless [of b]
 15.1594 -  zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2)
 15.1595 -done
 15.1596 -
 15.1597 -lemma zdiv_zmult_zmult1: "intify(c) \<noteq> #0 ==> (c$*a) zdiv (c$*b) = a zdiv b"
 15.1598 -apply (cut_tac b = "intify (b)" in zdiv_zmult_zmult1_raw)
 15.1599 -apply auto
 15.1600 -done
 15.1601 -
 15.1602 -lemma zdiv_zmult_zmult2: "intify(c) \<noteq> #0 ==> (a$*c) zdiv (b$*c) = a zdiv b"
 15.1603 -apply (drule zdiv_zmult_zmult1)
 15.1604 -apply (auto simp add: zmult_commute)
 15.1605 -done
 15.1606 -
 15.1607 -
 15.1608 -subsection{* Distribution of factors over "zmod" *}
 15.1609 -
 15.1610 -lemma zmod_zmult_zmult1_aux1:
 15.1611 -     "[| #0 $< b;  intify(c) \<noteq> #0 |]  
 15.1612 -      ==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
 15.1613 -apply (subst zmod_zmult2_eq)
 15.1614 -apply auto
 15.1615 -done
 15.1616 -
 15.1617 -lemma zmod_zmult_zmult1_aux2:
 15.1618 -     "[| b $< #0;  intify(c) \<noteq> #0 |]  
 15.1619 -      ==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
 15.1620 -apply (subgoal_tac " (c $* ($-a)) zmod (c $* ($-b)) = c $* (($-a) zmod ($-b))")
 15.1621 -apply (rule_tac [2] zmod_zmult_zmult1_aux1)
 15.1622 -apply auto
 15.1623 -done
 15.1624 -
 15.1625 -lemma zmod_zmult_zmult1_raw:
 15.1626 -     "[|b \<in> int; c \<in> int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
 15.1627 -apply (case_tac "b = #0")
 15.1628 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1629 -apply (case_tac "c = #0")
 15.1630 - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 15.1631 -apply (auto simp add: neq_iff_zless [of b]
 15.1632 -  zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2)
 15.1633 -done
 15.1634 -
 15.1635 -lemma zmod_zmult_zmult1: "(c$*a) zmod (c$*b) = c $* (a zmod b)"
 15.1636 -apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult_zmult1_raw)
 15.1637 -apply auto
 15.1638 -done
 15.1639 -
 15.1640 -lemma zmod_zmult_zmult2: "(a$*c) zmod (b$*c) = (a zmod b) $* c"
 15.1641 -apply (cut_tac c = "c" in zmod_zmult_zmult1)
 15.1642 -apply (auto simp add: zmult_commute)
 15.1643 -done
 15.1644 -
 15.1645 -
 15.1646 -(** Quotients of signs **)
 15.1647 -
 15.1648 -lemma zdiv_neg_pos_less0: "[| a $< #0;  #0 $< b |] ==> a zdiv b $< #0"
 15.1649 -apply (subgoal_tac "a zdiv b $<= #-1")
 15.1650 -apply (erule zle_zless_trans)
 15.1651 -apply (simp (no_asm))
 15.1652 -apply (rule zle_trans)
 15.1653 -apply (rule_tac a' = "#-1" in zdiv_mono1)
 15.1654 -apply (rule zless_add1_iff_zle [THEN iffD1])
 15.1655 -apply (simp (no_asm))
 15.1656 -apply (auto simp add: zdiv_minus1)
 15.1657 -done
 15.1658 -
 15.1659 -lemma zdiv_nonneg_neg_le0: "[| #0 $<= a;  b $< #0 |] ==> a zdiv b $<= #0"
 15.1660 -apply (drule zdiv_mono1_neg)
 15.1661 -apply auto
 15.1662 -done
 15.1663 -
 15.1664 -lemma pos_imp_zdiv_nonneg_iff: "#0 $< b ==> (#0 $<= a zdiv b) <-> (#0 $<= a)"
 15.1665 -apply auto
 15.1666 -apply (drule_tac [2] zdiv_mono1)
 15.1667 -apply (auto simp add: neq_iff_zless)
 15.1668 -apply (simp (no_asm_use) add: not_zless_iff_zle [THEN iff_sym])
 15.1669 -apply (blast intro: zdiv_neg_pos_less0)
 15.1670 -done
 15.1671 -
 15.1672 -lemma neg_imp_zdiv_nonneg_iff: "b $< #0 ==> (#0 $<= a zdiv b) <-> (a $<= #0)"
 15.1673 -apply (subst zdiv_zminus_zminus [symmetric])
 15.1674 -apply (rule iff_trans)
 15.1675 -apply (rule pos_imp_zdiv_nonneg_iff)
 15.1676 -apply auto
 15.1677 -done
 15.1678 -
 15.1679 -(*But not (a zdiv b $<= 0 iff a$<=0); consider a=1, b=2 when a zdiv b = 0.*)
 15.1680 -lemma pos_imp_zdiv_neg_iff: "#0 $< b ==> (a zdiv b $< #0) <-> (a $< #0)"
 15.1681 -apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym])
 15.1682 -apply (erule pos_imp_zdiv_nonneg_iff)
 15.1683 -done
 15.1684 -
 15.1685 -(*Again the law fails for $<=: consider a = -1, b = -2 when a zdiv b = 0*)
 15.1686 -lemma neg_imp_zdiv_neg_iff: "b $< #0 ==> (a zdiv b $< #0) <-> (#0 $< a)"
 15.1687 -apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym])
 15.1688 -apply (erule neg_imp_zdiv_nonneg_iff)
 15.1689 -done
 15.1690 -
 15.1691 -(*
 15.1692 - THESE REMAIN TO BE CONVERTED -- but aren't that useful!
 15.1693 -
 15.1694 - subsection{* Speeding up the division algorithm with shifting *}
 15.1695 -
 15.1696 - (** computing "zdiv" by shifting **)
 15.1697 -
 15.1698 - lemma pos_zdiv_mult_2: "#0 $<= a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a"
 15.1699 - apply (case_tac "a = #0")
 15.1700 - apply (subgoal_tac "#1 $<= a")
 15.1701 -  apply (arith_tac 2)
 15.1702 - apply (subgoal_tac "#1 $< a $* #2")
 15.1703 -  apply (arith_tac 2)
 15.1704 - apply (subgoal_tac "#2$* (#1 $+ b zmod a) $<= #2$*a")
 15.1705 -  apply (rule_tac [2] zmult_zle_mono2)
 15.1706 - apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound)
 15.1707 - apply (subst zdiv_zadd1_eq)
 15.1708 - apply (simp (no_asm_simp) add: zdiv_zmult_zmult2 zmod_zmult_zmult2 zdiv_pos_pos_trivial)
 15.1709 - apply (subst zdiv_pos_pos_trivial)
 15.1710 - apply (simp (no_asm_simp) add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ])
 15.1711 - apply (auto simp add: zmod_pos_pos_trivial)
 15.1712 - apply (subgoal_tac "#0 $<= b zmod a")
 15.1713 -  apply (asm_simp_tac (simpset () add: pos_mod_sign) 2)
 15.1714 - apply arith
 15.1715 - done
 15.1716 -
 15.1717 -
 15.1718 - lemma neg_zdiv_mult_2: "a $<= #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) <-> (b$+#1) zdiv a"
 15.1719 - apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zdiv (#2 $* ($-a)) <-> ($-b-#1) zdiv ($-a)")
 15.1720 - apply (rule_tac [2] pos_zdiv_mult_2)
 15.1721 - apply (auto simp add: zmult_zminus_right)
 15.1722 - apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))")
 15.1723 - apply (Simp_tac 2)
 15.1724 - apply (asm_full_simp_tac (HOL_ss add: zdiv_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric])
 15.1725 - done
 15.1726 -
 15.1727 -
 15.1728 - (*Not clear why this must be proved separately; probably integ_of causes
 15.1729 -   simplification problems*)
 15.1730 - lemma lemma: "~ #0 $<= x ==> x $<= #0"
 15.1731 - apply auto
 15.1732 - done
 15.1733 -
 15.1734 - lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) =  
 15.1735 -           (if ~b | #0 $<= integ_of w                    
 15.1736 -            then integ_of v zdiv (integ_of w)     
 15.1737 -            else (integ_of v $+ #1) zdiv (integ_of w))"
 15.1738 - apply (simp_tac (simpset_of Int.thy add: zadd_assoc integ_of_BIT)
 15.1739 - 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)
 15.1740 - done
 15.1741 -
 15.1742 - declare zdiv_integ_of_BIT [simp]
 15.1743 -
 15.1744 -
 15.1745 - (** computing "zmod" by shifting (proofs resemble those for "zdiv") **)
 15.1746 -
 15.1747 - lemma pos_zmod_mult_2: "#0 $<= a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)"
 15.1748 - apply (case_tac "a = #0")
 15.1749 - apply (subgoal_tac "#1 $<= a")
 15.1750 -  apply (arith_tac 2)
 15.1751 - apply (subgoal_tac "#1 $< a $* #2")
 15.1752 -  apply (arith_tac 2)
 15.1753 - apply (subgoal_tac "#2$* (#1 $+ b zmod a) $<= #2$*a")
 15.1754 -  apply (rule_tac [2] zmult_zle_mono2)
 15.1755 - apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound)
 15.1756 - apply (subst zmod_zadd1_eq)
 15.1757 - apply (simp (no_asm_simp) add: zmod_zmult_zmult2 zmod_pos_pos_trivial)
 15.1758 - apply (rule zmod_pos_pos_trivial)
 15.1759 - apply (simp (no_asm_simp) # add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ])
 15.1760 - apply (auto simp add: zmod_pos_pos_trivial)
 15.1761 - apply (subgoal_tac "#0 $<= b zmod a")
 15.1762 -  apply (asm_simp_tac (simpset () add: pos_mod_sign) 2)
 15.1763 - apply arith
 15.1764 - done
 15.1765 -
 15.1766 -
 15.1767 - lemma neg_zmod_mult_2: "a $<= #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1"
 15.1768 - apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zmod (#2$* ($-a)) = #1 $+ #2$* (($-b-#1) zmod ($-a))")
 15.1769 - apply (rule_tac [2] pos_zmod_mult_2)
 15.1770 - apply (auto simp add: zmult_zminus_right)
 15.1771 - apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))")
 15.1772 - apply (Simp_tac 2)
 15.1773 - apply (asm_full_simp_tac (HOL_ss add: zmod_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric])
 15.1774 - apply (dtac (zminus_equation [THEN iffD1, symmetric])
 15.1775 - apply auto
 15.1776 - done
 15.1777 -
 15.1778 - lemma zmod_integ_of_BIT: "integ_of (v BIT b) zmod integ_of (w BIT False) =  
 15.1779 -           (if b then  
 15.1780 -                 if #0 $<= integ_of w  
 15.1781 -                 then #2 $* (integ_of v zmod integ_of w) $+ #1     
 15.1782 -                 else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1   
 15.1783 -            else #2 $* (integ_of v zmod integ_of w))"
 15.1784 - apply (simp_tac (simpset_of Int.thy add: zadd_assoc integ_of_BIT)
 15.1785 - 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)
 15.1786 - done
 15.1787 -
 15.1788 - declare zmod_integ_of_BIT [simp]
 15.1789 -*)
 15.1790 -
 15.1791 -end
 15.1792 -
    16.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    16.2 +++ b/src/ZF/IntDiv_ZF.thy	Mon Feb 11 15:40:21 2008 +0100
    16.3 @@ -0,0 +1,1789 @@
    16.4 +(*  Title:      ZF/IntDiv.thy
    16.5 +    ID:         $Id$
    16.6 +    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    16.7 +    Copyright   1999  University of Cambridge
    16.8 +
    16.9 +Here is the division algorithm in ML:
   16.10 +
   16.11 +    fun posDivAlg (a,b) =
   16.12 +      if a<b then (0,a)
   16.13 +      else let val (q,r) = posDivAlg(a, 2*b)
   16.14 +	       in  if 0<=r-b then (2*q+1, r-b) else (2*q, r)
   16.15 +	   end
   16.16 +
   16.17 +    fun negDivAlg (a,b) =
   16.18 +      if 0<=a+b then (~1,a+b)
   16.19 +      else let val (q,r) = negDivAlg(a, 2*b)
   16.20 +	       in  if 0<=r-b then (2*q+1, r-b) else (2*q, r)
   16.21 +	   end;
   16.22 +
   16.23 +    fun negateSnd (q,r:int) = (q,~r);
   16.24 +
   16.25 +    fun divAlg (a,b) = if 0<=a then 
   16.26 +			  if b>0 then posDivAlg (a,b) 
   16.27 +			   else if a=0 then (0,0)
   16.28 +				else negateSnd (negDivAlg (~a,~b))
   16.29 +		       else 
   16.30 +			  if 0<b then negDivAlg (a,b)
   16.31 +			  else        negateSnd (posDivAlg (~a,~b));
   16.32 +
   16.33 +*)
   16.34 +
   16.35 +header{*The Division Operators Div and Mod*}
   16.36 +
   16.37 +theory IntDiv_ZF imports IntArith OrderArith begin
   16.38 +
   16.39 +definition
   16.40 +  quorem :: "[i,i] => o"  where
   16.41 +    "quorem == %<a,b> <q,r>.
   16.42 +                      a = b$*q $+ r &
   16.43 +                      (#0$<b & #0$<=r & r$<b | ~(#0$<b) & b$<r & r $<= #0)"
   16.44 +
   16.45 +definition
   16.46 +  adjust :: "[i,i] => i"  where
   16.47 +    "adjust(b) == %<q,r>. if #0 $<= r$-b then <#2$*q $+ #1,r$-b>
   16.48 +                          else <#2$*q,r>"
   16.49 +
   16.50 +
   16.51 +(** the division algorithm **)
   16.52 +
   16.53 +definition
   16.54 +  posDivAlg :: "i => i"  where
   16.55 +(*for the case a>=0, b>0*)
   16.56 +(*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*)
   16.57 +    "posDivAlg(ab) ==
   16.58 +       wfrec(measure(int*int, %<a,b>. nat_of (a $- b $+ #1)),
   16.59 +	     ab,
   16.60 +	     %<a,b> f. if (a$<b | b$<=#0) then <#0,a>
   16.61 +                       else adjust(b, f ` <a,#2$*b>))"
   16.62 +
   16.63 +
   16.64 +(*for the case a<0, b>0*)
   16.65 +definition
   16.66 +  negDivAlg :: "i => i"  where
   16.67 +(*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*)
   16.68 +    "negDivAlg(ab) ==
   16.69 +       wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)),
   16.70 +	     ab,
   16.71 +	     %<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b>
   16.72 +                       else adjust(b, f ` <a,#2$*b>))"
   16.73 +
   16.74 +(*for the general case b\<noteq>0*)
   16.75 +
   16.76 +definition
   16.77 +  negateSnd :: "i => i"  where
   16.78 +    "negateSnd == %<q,r>. <q, $-r>"
   16.79 +
   16.80 +  (*The full division algorithm considers all possible signs for a, b
   16.81 +    including the special case a=0, b<0, because negDivAlg requires a<0*)
   16.82 +definition
   16.83 +  divAlg :: "i => i"  where
   16.84 +    "divAlg ==
   16.85 +       %<a,b>. if #0 $<= a then
   16.86 +                  if #0 $<= b then posDivAlg (<a,b>)
   16.87 +                  else if a=#0 then <#0,#0>
   16.88 +                       else negateSnd (negDivAlg (<$-a,$-b>))
   16.89 +               else 
   16.90 +                  if #0$<b then negDivAlg (<a,b>)
   16.91 +                  else         negateSnd (posDivAlg (<$-a,$-b>))"
   16.92 +
   16.93 +definition
   16.94 +  zdiv  :: "[i,i]=>i"                    (infixl "zdiv" 70)  where
   16.95 +    "a zdiv b == fst (divAlg (<intify(a), intify(b)>))"
   16.96 +
   16.97 +definition
   16.98 +  zmod  :: "[i,i]=>i"                    (infixl "zmod" 70)  where
   16.99 +    "a zmod b == snd (divAlg (<intify(a), intify(b)>))"
  16.100 +
  16.101 +
  16.102 +(** Some basic laws by Sidi Ehmety (need linear arithmetic!) **)
  16.103 +
  16.104 +lemma zspos_add_zspos_imp_zspos: "[| #0 $< x;  #0 $< y |] ==> #0 $< x $+ y"
  16.105 +apply (rule_tac y = "y" in zless_trans)
  16.106 +apply (rule_tac [2] zdiff_zless_iff [THEN iffD1])
  16.107 +apply auto
  16.108 +done
  16.109 +
  16.110 +lemma zpos_add_zpos_imp_zpos: "[| #0 $<= x;  #0 $<= y |] ==> #0 $<= x $+ y"
  16.111 +apply (rule_tac y = "y" in zle_trans)
  16.112 +apply (rule_tac [2] zdiff_zle_iff [THEN iffD1])
  16.113 +apply auto
  16.114 +done
  16.115 +
  16.116 +lemma zneg_add_zneg_imp_zneg: "[| x $< #0;  y $< #0 |] ==> x $+ y $< #0"
  16.117 +apply (rule_tac y = "y" in zless_trans)
  16.118 +apply (rule zless_zdiff_iff [THEN iffD1])
  16.119 +apply auto
  16.120 +done
  16.121 +
  16.122 +(* this theorem is used below *)
  16.123 +lemma zneg_or_0_add_zneg_or_0_imp_zneg_or_0:
  16.124 +     "[| x $<= #0;  y $<= #0 |] ==> x $+ y $<= #0"
  16.125 +apply (rule_tac y = "y" in zle_trans)
  16.126 +apply (rule zle_zdiff_iff [THEN iffD1])
  16.127 +apply auto
  16.128 +done
  16.129 +
  16.130 +lemma zero_lt_zmagnitude: "[| #0 $< k; k \<in> int |] ==> 0 < zmagnitude(k)"
  16.131 +apply (drule zero_zless_imp_znegative_zminus)
  16.132 +apply (drule_tac [2] zneg_int_of)
  16.133 +apply (auto simp add: zminus_equation [of k])
  16.134 +apply (subgoal_tac "0 < zmagnitude ($# succ (n))")
  16.135 + apply simp
  16.136 +apply (simp only: zmagnitude_int_of)
  16.137 +apply simp
  16.138 +done
  16.139 +
  16.140 +
  16.141 +(*** Inequality lemmas involving $#succ(m) ***)
  16.142 +
  16.143 +lemma zless_add_succ_iff:
  16.144 +     "(w $< z $+ $# succ(m)) <-> (w $< z $+ $#m | intify(w) = z $+ $#m)"
  16.145 +apply (auto simp add: zless_iff_succ_zadd zadd_assoc int_of_add [symmetric])
  16.146 +apply (rule_tac [3] x = "0" in bexI)
  16.147 +apply (cut_tac m = "m" in int_succ_int_1)
  16.148 +apply (cut_tac m = "n" in int_succ_int_1)
  16.149 +apply simp
  16.150 +apply (erule natE)
  16.151 +apply auto
  16.152 +apply (rule_tac x = "succ (n) " in bexI)
  16.153 +apply auto
  16.154 +done
  16.155 +
  16.156 +lemma zadd_succ_lemma:
  16.157 +     "z \<in> int ==> (w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)"
  16.158 +apply (simp only: not_zless_iff_zle [THEN iff_sym] zless_add_succ_iff)
  16.159 +apply (auto intro: zle_anti_sym elim: zless_asym
  16.160 +            simp add: zless_imp_zle not_zless_iff_zle)
  16.161 +done
  16.162 +
  16.163 +lemma zadd_succ_zle_iff: "(w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)"
  16.164 +apply (cut_tac z = "intify (z)" in zadd_succ_lemma)
  16.165 +apply auto
  16.166 +done
  16.167 +
  16.168 +(** Inequality reasoning **)
  16.169 +
  16.170 +lemma zless_add1_iff_zle: "(w $< z $+ #1) <-> (w$<=z)"
  16.171 +apply (subgoal_tac "#1 = $# 1")
  16.172 +apply (simp only: zless_add_succ_iff zle_def)
  16.173 +apply auto
  16.174 +done
  16.175 +
  16.176 +lemma add1_zle_iff: "(w $+ #1 $<= z) <-> (w $< z)"
  16.177 +apply (subgoal_tac "#1 = $# 1")
  16.178 +apply (simp only: zadd_succ_zle_iff)
  16.179 +apply auto
  16.180 +done
  16.181 +
  16.182 +lemma add1_left_zle_iff: "(#1 $+ w $<= z) <-> (w $< z)"
  16.183 +apply (subst zadd_commute)
  16.184 +apply (rule add1_zle_iff)
  16.185 +done
  16.186 +
  16.187 +
  16.188 +(*** Monotonicity of Multiplication ***)
  16.189 +
  16.190 +lemma zmult_mono_lemma: "k \<in> nat ==> i $<= j ==> i $* $#k $<= j $* $#k"
  16.191 +apply (induct_tac "k")
  16.192 + prefer 2 apply (subst int_succ_int_1)
  16.193 +apply (simp_all (no_asm_simp) add: zadd_zmult_distrib2 zadd_zle_mono)
  16.194 +done
  16.195 +
  16.196 +lemma zmult_zle_mono1: "[| i $<= j;  #0 $<= k |] ==> i$*k $<= j$*k"
  16.197 +apply (subgoal_tac "i $* intify (k) $<= j $* intify (k) ")
  16.198 +apply (simp (no_asm_use))
  16.199 +apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst])
  16.200 +apply (rule_tac [3] zmult_mono_lemma)
  16.201 +apply auto
  16.202 +apply (simp add: znegative_iff_zless_0 not_zless_iff_zle [THEN iff_sym])
  16.203 +done
  16.204 +
  16.205 +lemma zmult_zle_mono1_neg: "[| i $<= j;  k $<= #0 |] ==> j$*k $<= i$*k"
  16.206 +apply (rule zminus_zle_zminus [THEN iffD1])
  16.207 +apply (simp del: zmult_zminus_right
  16.208 +            add: zmult_zminus_right [symmetric] zmult_zle_mono1 zle_zminus)
  16.209 +done
  16.210 +
  16.211 +lemma zmult_zle_mono2: "[| i $<= j;  #0 $<= k |] ==> k$*i $<= k$*j"
  16.212 +apply (drule zmult_zle_mono1)
  16.213 +apply (simp_all add: zmult_commute)
  16.214 +done
  16.215 +
  16.216 +lemma zmult_zle_mono2_neg: "[| i $<= j;  k $<= #0 |] ==> k$*j $<= k$*i"
  16.217 +apply (drule zmult_zle_mono1_neg)
  16.218 +apply (simp_all add: zmult_commute)
  16.219 +done
  16.220 +
  16.221 +(* $<= monotonicity, BOTH arguments*)
  16.222 +lemma zmult_zle_mono:
  16.223 +     "[| i $<= j;  k $<= l;  #0 $<= j;  #0 $<= k |] ==> i$*k $<= j$*l"
  16.224 +apply (erule zmult_zle_mono1 [THEN zle_trans])
  16.225 +apply assumption
  16.226 +apply (erule zmult_zle_mono2)
  16.227 +apply assumption
  16.228 +done
  16.229 +
  16.230 +
  16.231 +(** strict, in 1st argument; proof is by induction on k>0 **)
  16.232 +
  16.233 +lemma zmult_zless_mono2_lemma [rule_format]:
  16.234 +     "[| i$<j; k \<in> nat |] ==> 0<k --> $#k $* i $< $#k $* j"
  16.235 +apply (induct_tac "k")
  16.236 + prefer 2
  16.237 + apply (subst int_succ_int_1)
  16.238 + apply (erule natE)
  16.239 +apply (simp_all add: zadd_zmult_distrib zadd_zless_mono zle_def)
  16.240 +apply (frule nat_0_le)
  16.241 +apply (subgoal_tac "i $+ (i $+ $# xa $* i) $< j $+ (j $+ $# xa $* j) ")
  16.242 +apply (simp (no_asm_use))
  16.243 +apply (rule zadd_zless_mono)
  16.244 +apply (simp_all (no_asm_simp) add: zle_def)
  16.245 +done
  16.246 +
  16.247 +lemma zmult_zless_mono2: "[| i$<j;  #0 $< k |] ==> k$*i $< k$*j"
  16.248 +apply (subgoal_tac "intify (k) $* i $< intify (k) $* j")
  16.249 +apply (simp (no_asm_use))
  16.250 +apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst])
  16.251 +apply (rule_tac [3] zmult_zless_mono2_lemma)
  16.252 +apply auto
  16.253 +apply (simp add: znegative_iff_zless_0)
  16.254 +apply (drule zless_trans, assumption)
  16.255 +apply (auto simp add: zero_lt_zmagnitude)
  16.256 +done
  16.257 +
  16.258 +lemma zmult_zless_mono1: "[| i$<j;  #0 $< k |] ==> i$*k $< j$*k"
  16.259 +apply (drule zmult_zless_mono2)
  16.260 +apply (simp_all add: zmult_commute)
  16.261 +done
  16.262 +
  16.263 +(* < monotonicity, BOTH arguments*)
  16.264 +lemma zmult_zless_mono:
  16.265 +     "[| i $< j;  k $< l;  #0 $< j;  #0 $< k |] ==> i$*k $< j$*l"
  16.266 +apply (erule zmult_zless_mono1 [THEN zless_trans])
  16.267 +apply assumption
  16.268 +apply (erule zmult_zless_mono2)
  16.269 +apply assumption
  16.270 +done
  16.271 +
  16.272 +lemma zmult_zless_mono1_neg: "[| i $< j;  k $< #0 |] ==> j$*k $< i$*k"
  16.273 +apply (rule zminus_zless_zminus [THEN iffD1])
  16.274 +apply (simp del: zmult_zminus_right 
  16.275 +            add: zmult_zminus_right [symmetric] zmult_zless_mono1 zless_zminus)
  16.276 +done
  16.277 +
  16.278 +lemma zmult_zless_mono2_neg: "[| i $< j;  k $< #0 |] ==> k$*j $< k$*i"
  16.279 +apply (rule zminus_zless_zminus [THEN iffD1])
  16.280 +apply (simp del: zmult_zminus 
  16.281 +            add: zmult_zminus [symmetric] zmult_zless_mono2 zless_zminus)
  16.282 +done
  16.283 +
  16.284 +
  16.285 +(** Products of zeroes **)
  16.286 +
  16.287 +lemma zmult_eq_lemma:
  16.288 +     "[| m \<in> int; n \<in> int |] ==> (m = #0 | n = #0) <-> (m$*n = #0)"
  16.289 +apply (case_tac "m $< #0")
  16.290 +apply (auto simp add: not_zless_iff_zle zle_def neq_iff_zless)
  16.291 +apply (force dest: zmult_zless_mono1_neg zmult_zless_mono1)+
  16.292 +done
  16.293 +
  16.294 +lemma zmult_eq_0_iff [iff]: "(m$*n = #0) <-> (intify(m) = #0 | intify(n) = #0)"
  16.295 +apply (simp add: zmult_eq_lemma)
  16.296 +done
  16.297 +
  16.298 +
  16.299 +(** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =,
  16.300 +    but not (yet?) for k*m < n*k. **)
  16.301 +
  16.302 +lemma zmult_zless_lemma:
  16.303 +     "[| k \<in> int; m \<in> int; n \<in> int |]  
  16.304 +      ==> (m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
  16.305 +apply (case_tac "k = #0")
  16.306 +apply (auto simp add: neq_iff_zless zmult_zless_mono1 zmult_zless_mono1_neg)
  16.307 +apply (auto simp add: not_zless_iff_zle 
  16.308 +                      not_zle_iff_zless [THEN iff_sym, of "m$*k"] 
  16.309 +                      not_zle_iff_zless [THEN iff_sym, of m])
  16.310 +apply (auto elim: notE
  16.311 +            simp add: zless_imp_zle zmult_zle_mono1 zmult_zle_mono1_neg)
  16.312 +done
  16.313 +
  16.314 +lemma zmult_zless_cancel2:
  16.315 +     "(m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
  16.316 +apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)" 
  16.317 +       in zmult_zless_lemma)
  16.318 +apply auto
  16.319 +done
  16.320 +
  16.321 +lemma zmult_zless_cancel1:
  16.322 +     "(k$*m $< k$*n) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
  16.323 +by (simp add: zmult_commute [of k] zmult_zless_cancel2)
  16.324 +
  16.325 +lemma zmult_zle_cancel2:
  16.326 +     "(m$*k $<= n$*k) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))"
  16.327 +by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel2)
  16.328 +
  16.329 +lemma zmult_zle_cancel1:
  16.330 +     "(k$*m $<= k$*n) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))"
  16.331 +by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel1)
  16.332 +
  16.333 +lemma int_eq_iff_zle: "[| m \<in> int; n \<in> int |] ==> m=n <-> (m $<= n & n $<= m)"
  16.334 +apply (blast intro: zle_refl zle_anti_sym)
  16.335 +done
  16.336 +
  16.337 +lemma zmult_cancel2_lemma:
  16.338 +     "[| k \<in> int; m \<in> int; n \<in> int |] ==> (m$*k = n$*k) <-> (k=#0 | m=n)"
  16.339 +apply (simp add: int_eq_iff_zle [of "m$*k"] int_eq_iff_zle [of m])
  16.340 +apply (auto simp add: zmult_zle_cancel2 neq_iff_zless)
  16.341 +done
  16.342 +
  16.343 +lemma zmult_cancel2 [simp]:
  16.344 +     "(m$*k = n$*k) <-> (intify(k) = #0 | intify(m) = intify(n))"
  16.345 +apply (rule iff_trans)
  16.346 +apply (rule_tac [2] zmult_cancel2_lemma)
  16.347 +apply auto
  16.348 +done
  16.349 +
  16.350 +lemma zmult_cancel1 [simp]:
  16.351 +     "(k$*m = k$*n) <-> (intify(k) = #0 | intify(m) = intify(n))"
  16.352 +by (simp add: zmult_commute [of k] zmult_cancel2)
  16.353 +
  16.354 +
  16.355 +subsection{* Uniqueness and monotonicity of quotients and remainders *}
  16.356 +
  16.357 +lemma unique_quotient_lemma:
  16.358 +     "[| b$*q' $+ r' $<= b$*q $+ r;  #0 $<= r';  #0 $< b;  r $< b |]  
  16.359 +      ==> q' $<= q"
  16.360 +apply (subgoal_tac "r' $+ b $* (q'$-q) $<= r")
  16.361 + prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_ac zcompare_rls)
  16.362 +apply (subgoal_tac "#0 $< b $* (#1 $+ q $- q') ")
  16.363 + prefer 2
  16.364 + apply (erule zle_zless_trans)
  16.365 + apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls)
  16.366 + apply (erule zle_zless_trans)
  16.367 + apply (simp add: ); 
  16.368 +apply (subgoal_tac "b $* q' $< b $* (#1 $+ q)")
  16.369 + prefer 2 
  16.370 + apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls)
  16.371 +apply (auto elim: zless_asym
  16.372 +        simp add: zmult_zless_cancel1 zless_add1_iff_zle zadd_ac zcompare_rls)
  16.373 +done
  16.374 +
  16.375 +lemma unique_quotient_lemma_neg:
  16.376 +     "[| b$*q' $+ r' $<= b$*q $+ r;  r $<= #0;  b $< #0;  b $< r' |]  
  16.377 +      ==> q $<= q'"
  16.378 +apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r" 
  16.379 +       in unique_quotient_lemma)
  16.380 +apply (auto simp del: zminus_zadd_distrib
  16.381 +            simp add: zminus_zadd_distrib [symmetric] zle_zminus zless_zminus)
  16.382 +done
  16.383 +
  16.384 +
  16.385 +lemma unique_quotient:
  16.386 +     "[| quorem (<a,b>, <q,r>);  quorem (<a,b>, <q',r'>);  b \<in> int; b ~= #0;  
  16.387 +         q \<in> int; q' \<in> int |] ==> q = q'"
  16.388 +apply (simp add: split_ifs quorem_def neq_iff_zless)
  16.389 +apply safe
  16.390 +apply simp_all
  16.391 +apply (blast intro: zle_anti_sym
  16.392 +             dest: zle_eq_refl [THEN unique_quotient_lemma] 
  16.393 +                   zle_eq_refl [THEN unique_quotient_lemma_neg] sym)+
  16.394 +done
  16.395 +
  16.396 +lemma unique_remainder:
  16.397 +     "[| quorem (<a,b>, <q,r>);  quorem (<a,b>, <q',r'>);  b \<in> int; b ~= #0;  
  16.398 +         q \<in> int; q' \<in> int;  
  16.399 +         r \<in> int; r' \<in> int |] ==> r = r'"
  16.400 +apply (subgoal_tac "q = q'")
  16.401 + prefer 2 apply (blast intro: unique_quotient)
  16.402 +apply (simp add: quorem_def)
  16.403 +done
  16.404 +
  16.405 +
  16.406 +subsection{*Correctness of posDivAlg, 
  16.407 +           the Division Algorithm for @{text "a\<ge>0"} and @{text "b>0"} *}
  16.408 +
  16.409 +lemma adjust_eq [simp]:
  16.410 +     "adjust(b, <q,r>) = (let diff = r$-b in  
  16.411 +                          if #0 $<= diff then <#2$*q $+ #1,diff>   
  16.412 +                                         else <#2$*q,r>)"
  16.413 +by (simp add: Let_def adjust_def)
  16.414 +
  16.415 +
  16.416 +lemma posDivAlg_termination:
  16.417 +     "[| #0 $< b; ~ a $< b |]    
  16.418 +      ==> nat_of(a $- #2 $\<times> b $+ #1) < nat_of(a $- b $+ #1)"
  16.419 +apply (simp (no_asm) add: zless_nat_conj)
  16.420 +apply (simp add: not_zless_iff_zle zless_add1_iff_zle zcompare_rls)
  16.421 +done
  16.422 +
  16.423 +lemmas posDivAlg_unfold = def_wfrec [OF posDivAlg_def wf_measure]
  16.424 +
  16.425 +lemma posDivAlg_eqn:
  16.426 +     "[| #0 $< b; a \<in> int; b \<in> int |] ==>  
  16.427 +      posDivAlg(<a,b>) =       
  16.428 +       (if a$<b then <#0,a> else adjust(b, posDivAlg (<a, #2$*b>)))"
  16.429 +apply (rule posDivAlg_unfold [THEN trans])
  16.430 +apply (simp add: vimage_iff not_zless_iff_zle [THEN iff_sym])
  16.431 +apply (blast intro: posDivAlg_termination)
  16.432 +done
  16.433 +
  16.434 +lemma posDivAlg_induct_lemma [rule_format]:
  16.435 +  assumes prem:
  16.436 +        "!!a b. [| a \<in> int; b \<in> int;  
  16.437 +                   ~ (a $< b | b $<= #0) --> P(<a, #2 $* b>) |] ==> P(<a,b>)"
  16.438 +  shows "<u,v> \<in> int*int --> P(<u,v>)"
  16.439 +apply (rule_tac a = "<u,v>" in wf_induct)
  16.440 +apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)" 
  16.441 +       in wf_measure)
  16.442 +apply clarify
  16.443 +apply (rule prem)
  16.444 +apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec)
  16.445 +apply auto
  16.446 +apply (simp add: not_zle_iff_zless posDivAlg_termination)
  16.447 +done
  16.448 +
  16.449 +
  16.450 +lemma posDivAlg_induct [consumes 2]:
  16.451 +  assumes u_int: "u \<in> int"
  16.452 +      and v_int: "v \<in> int"
  16.453 +      and ih: "!!a b. [| a \<in> int; b \<in> int;
  16.454 +                     ~ (a $< b | b $<= #0) --> P(a, #2 $* b) |] ==> P(a,b)"
  16.455 +  shows "P(u,v)"
  16.456 +apply (subgoal_tac "(%<x,y>. P (x,y)) (<u,v>)")
  16.457 +apply simp
  16.458 +apply (rule posDivAlg_induct_lemma)
  16.459 +apply (simp (no_asm_use))
  16.460 +apply (rule ih)
  16.461 +apply (auto simp add: u_int v_int)
  16.462 +done
  16.463 +
  16.464 +(*FIXME: use intify in integ_of so that we always have integ_of w \<in> int.
  16.465 +    then this rewrite can work for ALL constants!!*)
  16.466 +lemma intify_eq_0_iff_zle: "intify(m) = #0 <-> (m $<= #0 & #0 $<= m)"
  16.467 +apply (simp (no_asm) add: int_eq_iff_zle)
  16.468 +done
  16.469 +
  16.470 +
  16.471 +subsection{* Some convenient biconditionals for products of signs *}
  16.472 +
  16.473 +lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j"
  16.474 +apply (drule zmult_zless_mono1)
  16.475 +apply auto
  16.476 +done
  16.477 +
  16.478 +lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j"
  16.479 +apply (drule zmult_zless_mono1_neg)
  16.480 +apply auto
  16.481 +done
  16.482 +
  16.483 +lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0"
  16.484 +apply (drule zmult_zless_mono1_neg)
  16.485 +apply auto
  16.486 +done
  16.487 +
  16.488 +(** Inequality reasoning **)
  16.489 +
  16.490 +lemma int_0_less_lemma:
  16.491 +     "[| x \<in> int; y \<in> int |]  
  16.492 +      ==> (#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)"
  16.493 +apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg)
  16.494 +apply (rule ccontr) 
  16.495 +apply (rule_tac [2] ccontr) 
  16.496 +apply (auto simp add: zle_def not_zless_iff_zle)
  16.497 +apply (erule_tac P = "#0$< x$* y" in rev_mp)
  16.498 +apply (erule_tac [2] P = "#0$< x$* y" in rev_mp)
  16.499 +apply (drule zmult_pos_neg, assumption) 
  16.500 + prefer 2
  16.501 + apply (drule zmult_pos_neg, assumption) 
  16.502 +apply (auto dest: zless_not_sym simp add: zmult_commute)
  16.503 +done
  16.504 +
  16.505 +lemma int_0_less_mult_iff:
  16.506 +     "(#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)"
  16.507 +apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_less_lemma)
  16.508 +apply auto
  16.509 +done
  16.510 +
  16.511 +lemma int_0_le_lemma:
  16.512 +     "[| x \<in> int; y \<in> int |]  
  16.513 +      ==> (#0 $<= x $* y) <-> (#0 $<= x & #0 $<= y | x $<= #0 & y $<= #0)"
  16.514 +by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff)
  16.515 +
  16.516 +lemma int_0_le_mult_iff:
  16.517 +     "(#0 $<= x $* y) <-> ((#0 $<= x & #0 $<= y) | (x $<= #0 & y $<= #0))"
  16.518 +apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_le_lemma)
  16.519 +apply auto
  16.520 +done
  16.521 +
  16.522 +lemma zmult_less_0_iff:
  16.523 +     "(x $* y $< #0) <-> (#0 $< x & y $< #0 | x $< #0 & #0 $< y)"
  16.524 +apply (auto simp add: int_0_le_mult_iff not_zle_iff_zless [THEN iff_sym])
  16.525 +apply (auto dest: zless_not_sym simp add: not_zle_iff_zless)
  16.526 +done
  16.527 +
  16.528 +lemma zmult_le_0_iff:
  16.529 +     "(x $* y $<= #0) <-> (#0 $<= x & y $<= #0 | x $<= #0 & #0 $<= y)"
  16.530 +by (auto dest: zless_not_sym
  16.531 +         simp add: int_0_less_mult_iff not_zless_iff_zle [THEN iff_sym])
  16.532 +
  16.533 +
  16.534 +(*Typechecking for posDivAlg*)
  16.535 +lemma posDivAlg_type [rule_format]:
  16.536 +     "[| a \<in> int; b \<in> int |] ==> posDivAlg(<a,b>) \<in> int * int"
  16.537 +apply (rule_tac u = "a" and v = "b" in posDivAlg_induct)
  16.538 +apply assumption+
  16.539 +apply (case_tac "#0 $< ba")
  16.540 + apply (simp add: posDivAlg_eqn adjust_def integ_of_type 
  16.541 +             split add: split_if_asm)
  16.542 + apply clarify
  16.543 + apply (simp add: int_0_less_mult_iff not_zle_iff_zless)
  16.544 +apply (simp add: not_zless_iff_zle)
  16.545 +apply (subst posDivAlg_unfold)
  16.546 +apply simp
  16.547 +done
  16.548 +
  16.549 +(*Correctness of posDivAlg: it computes quotients correctly*)
  16.550 +lemma posDivAlg_correct [rule_format]:
  16.551 +     "[| a \<in> int; b \<in> int |]  
  16.552 +      ==> #0 $<= a --> #0 $< b --> quorem (<a,b>, posDivAlg(<a,b>))"
  16.553 +apply (rule_tac u = "a" and v = "b" in posDivAlg_induct)
  16.554 +apply auto
  16.555 +   apply (simp_all add: quorem_def)
  16.556 +   txt{*base case: a<b*}
  16.557 +   apply (simp add: posDivAlg_eqn)
  16.558 +  apply (simp add: not_zless_iff_zle [THEN iff_sym])
  16.559 + apply (simp add: int_0_less_mult_iff)
  16.560 +txt{*main argument*}
  16.561 +apply (subst posDivAlg_eqn)
  16.562 +apply (simp_all (no_asm_simp))
  16.563 +apply (erule splitE)
  16.564 +apply (rule posDivAlg_type)
  16.565 +apply (simp_all add: int_0_less_mult_iff)
  16.566 +apply (auto simp add: zadd_zmult_distrib2 Let_def)
  16.567 +txt{*now just linear arithmetic*}
  16.568 +apply (simp add: not_zle_iff_zless zdiff_zless_iff)
  16.569 +done
  16.570 +
  16.571 +
  16.572 +subsection{*Correctness of negDivAlg, the division algorithm for a<0 and b>0*}
  16.573 +
  16.574 +lemma negDivAlg_termination:
  16.575 +     "[| #0 $< b; a $+ b $< #0 |] 
  16.576 +      ==> nat_of($- a $- #2 $* b) < nat_of($- a $- b)"
  16.577 +apply (simp (no_asm) add: zless_nat_conj)
  16.578 +apply (simp add: zcompare_rls not_zle_iff_zless zless_zdiff_iff [THEN iff_sym]
  16.579 +                 zless_zminus)
  16.580 +done
  16.581 +
  16.582 +lemmas negDivAlg_unfold = def_wfrec [OF negDivAlg_def wf_measure]
  16.583 +
  16.584 +lemma negDivAlg_eqn:
  16.585 +     "[| #0 $< b; a : int; b : int |] ==>  
  16.586 +      negDivAlg(<a,b>) =       
  16.587 +       (if #0 $<= a$+b then <#-1,a$+b>  
  16.588 +                       else adjust(b, negDivAlg (<a, #2$*b>)))"
  16.589 +apply (rule negDivAlg_unfold [THEN trans])
  16.590 +apply (simp (no_asm_simp) add: vimage_iff not_zless_iff_zle [THEN iff_sym])
  16.591 +apply (blast intro: negDivAlg_termination)
  16.592 +done
  16.593 +
  16.594 +lemma negDivAlg_induct_lemma [rule_format]:
  16.595 +  assumes prem:
  16.596 +        "!!a b. [| a \<in> int; b \<in> int;  
  16.597 +                   ~ (#0 $<= a $+ b | b $<= #0) --> P(<a, #2 $* b>) |]  
  16.598 +                ==> P(<a,b>)"
  16.599 +  shows "<u,v> \<in> int*int --> P(<u,v>)"
  16.600 +apply (rule_tac a = "<u,v>" in wf_induct)
  16.601 +apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)" 
  16.602 +       in wf_measure)
  16.603 +apply clarify
  16.604 +apply (rule prem)
  16.605 +apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec)
  16.606 +apply auto
  16.607 +apply (simp add: not_zle_iff_zless negDivAlg_termination)
  16.608 +done
  16.609 +
  16.610 +lemma negDivAlg_induct [consumes 2]:
  16.611 +  assumes u_int: "u \<in> int"
  16.612 +      and v_int: "v \<in> int"
  16.613 +      and ih: "!!a b. [| a \<in> int; b \<in> int;  
  16.614 +                         ~ (#0 $<= a $+ b | b $<= #0) --> P(a, #2 $* b) |]  
  16.615 +                      ==> P(a,b)"
  16.616 +  shows "P(u,v)"
  16.617 +apply (subgoal_tac " (%<x,y>. P (x,y)) (<u,v>)")
  16.618 +apply simp
  16.619 +apply (rule negDivAlg_induct_lemma)
  16.620 +apply (simp (no_asm_use))
  16.621 +apply (rule ih)
  16.622 +apply (auto simp add: u_int v_int)
  16.623 +done
  16.624 +
  16.625 +
  16.626 +(*Typechecking for negDivAlg*)
  16.627 +lemma negDivAlg_type:
  16.628 +     "[| a \<in> int; b \<in> int |] ==> negDivAlg(<a,b>) \<in> int * int"
  16.629 +apply (rule_tac u = "a" and v = "b" in negDivAlg_induct)
  16.630 +apply assumption+
  16.631 +apply (case_tac "#0 $< ba")
  16.632 + apply (simp add: negDivAlg_eqn adjust_def integ_of_type 
  16.633 +             split add: split_if_asm)
  16.634 + apply clarify
  16.635 + apply (simp add: int_0_less_mult_iff not_zle_iff_zless)
  16.636 +apply (simp add: not_zless_iff_zle)
  16.637 +apply (subst negDivAlg_unfold)
  16.638 +apply simp
  16.639 +done
  16.640 +
  16.641 +
  16.642 +(*Correctness of negDivAlg: it computes quotients correctly
  16.643 +  It doesn't work if a=0 because the 0/b=0 rather than -1*)
  16.644 +lemma negDivAlg_correct [rule_format]:
  16.645 +     "[| a \<in> int; b \<in> int |]  
  16.646 +      ==> a $< #0 --> #0 $< b --> quorem (<a,b>, negDivAlg(<a,b>))"
  16.647 +apply (rule_tac u = "a" and v = "b" in negDivAlg_induct)
  16.648 +  apply auto
  16.649 +   apply (simp_all add: quorem_def)
  16.650 +   txt{*base case: @{term "0$<=a$+b"}*}
  16.651 +   apply (simp add: negDivAlg_eqn)
  16.652 +  apply (simp add: not_zless_iff_zle [THEN iff_sym])
  16.653 + apply (simp add: int_0_less_mult_iff)
  16.654 +txt{*main argument*}
  16.655 +apply (subst negDivAlg_eqn)
  16.656 +apply (simp_all (no_asm_simp))
  16.657 +apply (erule splitE)
  16.658 +apply (rule negDivAlg_type)
  16.659 +apply (simp_all add: int_0_less_mult_iff)
  16.660 +apply (auto simp add: zadd_zmult_distrib2 Let_def)
  16.661 +txt{*now just linear arithmetic*}
  16.662 +apply (simp add: not_zle_iff_zless zdiff_zless_iff)
  16.663 +done
  16.664 +
  16.665 +
  16.666 +subsection{* Existence shown by proving the division algorithm to be correct *}
  16.667 +
  16.668 +(*the case a=0*)
  16.669 +lemma quorem_0: "[|b \<noteq> #0;  b \<in> int|] ==> quorem (<#0,b>, <#0,#0>)"
  16.670 +by (force simp add: quorem_def neq_iff_zless)
  16.671 +
  16.672 +lemma posDivAlg_zero_divisor: "posDivAlg(<a,#0>) = <#0,a>"
  16.673 +apply (subst posDivAlg_unfold)
  16.674 +apply simp
  16.675 +done
  16.676 +
  16.677 +lemma posDivAlg_0 [simp]: "posDivAlg (<#0,b>) = <#0,#0>"
  16.678 +apply (subst posDivAlg_unfold)
  16.679 +apply (simp add: not_zle_iff_zless)
  16.680 +done
  16.681 +
  16.682 +
  16.683 +(*Needed below.  Actually it's an equivalence.*)
  16.684 +lemma linear_arith_lemma: "~ (#0 $<= #-1 $+ b) ==> (b $<= #0)"
  16.685 +apply (simp add: not_zle_iff_zless)
  16.686 +apply (drule zminus_zless_zminus [THEN iffD2])
  16.687 +apply (simp add: zadd_commute zless_add1_iff_zle zle_zminus)
  16.688 +done
  16.689 +
  16.690 +lemma negDivAlg_minus1 [simp]: "negDivAlg (<#-1,b>) = <#-1, b$-#1>"
  16.691 +apply (subst negDivAlg_unfold)
  16.692 +apply (simp add: linear_arith_lemma integ_of_type vimage_iff)
  16.693 +done
  16.694 +
  16.695 +lemma negateSnd_eq [simp]: "negateSnd (<q,r>) = <q, $-r>"
  16.696 +apply (unfold negateSnd_def)
  16.697 +apply auto
  16.698 +done
  16.699 +
  16.700 +lemma negateSnd_type: "qr \<in> int * int ==> negateSnd (qr) \<in> int * int"
  16.701 +apply (unfold negateSnd_def)
  16.702 +apply auto
  16.703 +done
  16.704 +
  16.705 +lemma quorem_neg:
  16.706 +     "[|quorem (<$-a,$-b>, qr);  a \<in> int;  b \<in> int;  qr \<in> int * int|]   
  16.707 +      ==> quorem (<a,b>, negateSnd(qr))"
  16.708 +apply clarify
  16.709 +apply (auto elim: zless_asym simp add: quorem_def zless_zminus)
  16.710 +txt{*linear arithmetic from here on*}
  16.711 +apply (simp_all add: zminus_equation [of a] zminus_zless)
  16.712 +apply (cut_tac [2] z = "b" and w = "#0" in zless_linear)
  16.713 +apply (cut_tac [1] z = "b" and w = "#0" in zless_linear)
  16.714 +apply auto
  16.715 +apply (blast dest: zle_zless_trans)+
  16.716 +done
  16.717 +
  16.718 +lemma divAlg_correct:
  16.719 +     "[|b \<noteq> #0;  a \<in> int;  b \<in> int|] ==> quorem (<a,b>, divAlg(<a,b>))"
  16.720 +apply (auto simp add: quorem_0 divAlg_def)
  16.721 +apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct
  16.722 +                    posDivAlg_type negDivAlg_type) 
  16.723 +apply (auto simp add: quorem_def neq_iff_zless)
  16.724 +txt{*linear arithmetic from here on*}
  16.725 +apply (auto simp add: zle_def)
  16.726 +done
  16.727 +
  16.728 +lemma divAlg_type: "[|a \<in> int;  b \<in> int|] ==> divAlg(<a,b>) \<in> int * int"
  16.729 +apply (auto simp add: divAlg_def)
  16.730 +apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type)
  16.731 +done
  16.732 +
  16.733 +
  16.734 +(** intify cancellation **)
  16.735 +
  16.736 +lemma zdiv_intify1 [simp]: "intify(x) zdiv y = x zdiv y"
  16.737 +apply (simp (no_asm) add: zdiv_def)
  16.738 +done
  16.739 +
  16.740 +lemma zdiv_intify2 [simp]: "x zdiv intify(y) = x zdiv y"
  16.741 +apply (simp (no_asm) add: zdiv_def)
  16.742 +done
  16.743 +
  16.744 +lemma zdiv_type [iff,TC]: "z zdiv w \<in> int"
  16.745 +apply (unfold zdiv_def)
  16.746 +apply (blast intro: fst_type divAlg_type)
  16.747 +done
  16.748 +
  16.749 +lemma zmod_intify1 [simp]: "intify(x) zmod y = x zmod y"
  16.750 +apply (simp (no_asm) add: zmod_def)
  16.751 +done
  16.752 +
  16.753 +lemma zmod_intify2 [simp]: "x zmod intify(y) = x zmod y"
  16.754 +apply (simp (no_asm) add: zmod_def)
  16.755 +done
  16.756 +
  16.757 +lemma zmod_type [iff,TC]: "z zmod w \<in> int"
  16.758 +apply (unfold zmod_def)
  16.759 +apply (rule snd_type)
  16.760 +apply (blast intro: divAlg_type)
  16.761 +done
  16.762 +
  16.763 +
  16.764 +(** Arbitrary definitions for division by zero.  Useful to simplify 
  16.765 +    certain equations **)
  16.766 +
  16.767 +lemma DIVISION_BY_ZERO_ZDIV: "a zdiv #0 = #0"
  16.768 +apply (simp (no_asm) add: zdiv_def divAlg_def posDivAlg_zero_divisor)
  16.769 +done  (*NOT for adding to default simpset*)
  16.770 +
  16.771 +lemma DIVISION_BY_ZERO_ZMOD: "a zmod #0 = intify(a)"
  16.772 +apply (simp (no_asm) add: zmod_def divAlg_def posDivAlg_zero_divisor)
  16.773 +done  (*NOT for adding to default simpset*)
  16.774 +
  16.775 +
  16.776 +
  16.777 +(** Basic laws about division and remainder **)
  16.778 +
  16.779 +lemma raw_zmod_zdiv_equality:
  16.780 +     "[| a \<in> int; b \<in> int |] ==> a = b $* (a zdiv b) $+ (a zmod b)"
  16.781 +apply (case_tac "b = #0")
  16.782 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
  16.783 +apply (cut_tac a = "a" and b = "b" in divAlg_correct)
  16.784 +apply (auto simp add: quorem_def zdiv_def zmod_def split_def)
  16.785 +done
  16.786 +
  16.787 +lemma zmod_zdiv_equality: "intify(a) = b $* (a zdiv b) $+ (a zmod b)"
  16.788 +apply (rule trans)
  16.789 +apply (rule_tac b = "intify (b)" in raw_zmod_zdiv_equality)
  16.790 +apply auto
  16.791 +done
  16.792 +
  16.793 +lemma pos_mod: "#0 $< b ==> #0 $<= a zmod b & a zmod b $< b"
  16.794 +apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct)
  16.795 +apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def)
  16.796 +apply (blast dest: zle_zless_trans)+
  16.797 +done
  16.798 +
  16.799 +lemmas pos_mod_sign = pos_mod [THEN conjunct1, standard]
  16.800 +and    pos_mod_bound = pos_mod [THEN conjunct2, standard]
  16.801 +
  16.802 +lemma neg_mod: "b $< #0 ==> a zmod b $<= #0 & b $< a zmod b"
  16.803 +apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct)
  16.804 +apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def)
  16.805 +apply (blast dest: zle_zless_trans)
  16.806 +apply (blast dest: zless_trans)+
  16.807 +done
  16.808 +
  16.809 +lemmas neg_mod_sign = neg_mod [THEN conjunct1, standard]
  16.810 +and    neg_mod_bound = neg_mod [THEN conjunct2, standard]
  16.811 +
  16.812 +
  16.813 +(** proving general properties of zdiv and zmod **)
  16.814 +
  16.815 +lemma quorem_div_mod:
  16.816 +     "[|b \<noteq> #0;  a \<in> int;  b \<in> int |]  
  16.817 +      ==> quorem (<a,b>, <a zdiv b, a zmod b>)"
  16.818 +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
  16.819 +apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound 
  16.820 +                      neg_mod_sign neg_mod_bound)
  16.821 +done
  16.822 +
  16.823 +(*Surely quorem(<a,b>,<q,r>) implies a \<in> int, but it doesn't matter*)
  16.824 +lemma quorem_div:
  16.825 +     "[| quorem(<a,b>,<q,r>);  b \<noteq> #0;  a \<in> int;  b \<in> int;  q \<in> int |]  
  16.826 +      ==> a zdiv b = q"
  16.827 +by (blast intro: quorem_div_mod [THEN unique_quotient])
  16.828 +
  16.829 +lemma quorem_mod:
  16.830 +     "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int |] 
  16.831 +      ==> a zmod b = r"
  16.832 +by (blast intro: quorem_div_mod [THEN unique_remainder])
  16.833 +
  16.834 +lemma zdiv_pos_pos_trivial_raw:
  16.835 +     "[| a \<in> int;  b \<in> int;  #0 $<= a;  a $< b |] ==> a zdiv b = #0"
  16.836 +apply (rule quorem_div)
  16.837 +apply (auto simp add: quorem_def)
  16.838 +(*linear arithmetic*)
  16.839 +apply (blast dest: zle_zless_trans)+
  16.840 +done
  16.841 +
  16.842 +lemma zdiv_pos_pos_trivial: "[| #0 $<= a;  a $< b |] ==> a zdiv b = #0"
  16.843 +apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  16.844 +       in zdiv_pos_pos_trivial_raw)
  16.845 +apply auto
  16.846 +done
  16.847 +
  16.848 +lemma zdiv_neg_neg_trivial_raw:
  16.849 +     "[| a \<in> int;  b \<in> int;  a $<= #0;  b $< a |] ==> a zdiv b = #0"
  16.850 +apply (rule_tac r = "a" in quorem_div)
  16.851 +apply (auto simp add: quorem_def)
  16.852 +(*linear arithmetic*)
  16.853 +apply (blast dest: zle_zless_trans zless_trans)+
  16.854 +done
  16.855 +
  16.856 +lemma zdiv_neg_neg_trivial: "[| a $<= #0;  b $< a |] ==> a zdiv b = #0"
  16.857 +apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  16.858 +       in zdiv_neg_neg_trivial_raw)
  16.859 +apply auto
  16.860 +done
  16.861 +
  16.862 +lemma zadd_le_0_lemma: "[| a$+b $<= #0;  #0 $< a;  #0 $< b |] ==> False"
  16.863 +apply (drule_tac z' = "#0" and z = "b" in zadd_zless_mono)
  16.864 +apply (auto simp add: zle_def)
  16.865 +apply (blast dest: zless_trans)
  16.866 +done
  16.867 +
  16.868 +lemma zdiv_pos_neg_trivial_raw:
  16.869 +     "[| a \<in> int;  b \<in> int;  #0 $< a;  a$+b $<= #0 |] ==> a zdiv b = #-1"
  16.870 +apply (rule_tac r = "a $+ b" in quorem_div)
  16.871 +apply (auto simp add: quorem_def)
  16.872 +(*linear arithmetic*)
  16.873 +apply (blast dest: zadd_le_0_lemma zle_zless_trans)+
  16.874 +done
  16.875 +
  16.876 +lemma zdiv_pos_neg_trivial: "[| #0 $< a;  a$+b $<= #0 |] ==> a zdiv b = #-1"
  16.877 +apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  16.878 +       in zdiv_pos_neg_trivial_raw)
  16.879 +apply auto
  16.880 +done
  16.881 +
  16.882 +(*There is no zdiv_neg_pos_trivial because  #0 zdiv b = #0 would supersede it*)
  16.883 +
  16.884 +
  16.885 +lemma zmod_pos_pos_trivial_raw:
  16.886 +     "[| a \<in> int;  b \<in> int;  #0 $<= a;  a $< b |] ==> a zmod b = a"
  16.887 +apply (rule_tac q = "#0" in quorem_mod)
  16.888 +apply (auto simp add: quorem_def)
  16.889 +(*linear arithmetic*)
  16.890 +apply (blast dest: zle_zless_trans)+
  16.891 +done
  16.892 +
  16.893 +lemma zmod_pos_pos_trivial: "[| #0 $<= a;  a $< b |] ==> a zmod b = intify(a)"
  16.894 +apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  16.895 +       in zmod_pos_pos_trivial_raw)
  16.896 +apply auto
  16.897 +done
  16.898 +
  16.899 +lemma zmod_neg_neg_trivial_raw:
  16.900 +     "[| a \<in> int;  b \<in> int;  a $<= #0;  b $< a |] ==> a zmod b = a"
  16.901 +apply (rule_tac q = "#0" in quorem_mod)
  16.902 +apply (auto simp add: quorem_def)
  16.903 +(*linear arithmetic*)
  16.904 +apply (blast dest: zle_zless_trans zless_trans)+
  16.905 +done
  16.906 +
  16.907 +lemma zmod_neg_neg_trivial: "[| a $<= #0;  b $< a |] ==> a zmod b = intify(a)"
  16.908 +apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  16.909 +       in zmod_neg_neg_trivial_raw)
  16.910 +apply auto
  16.911 +done
  16.912 +
  16.913 +lemma zmod_pos_neg_trivial_raw:
  16.914 +     "[| a \<in> int;  b \<in> int;  #0 $< a;  a$+b $<= #0 |] ==> a zmod b = a$+b"
  16.915 +apply (rule_tac q = "#-1" in quorem_mod)
  16.916 +apply (auto simp add: quorem_def)
  16.917 +(*linear arithmetic*)
  16.918 +apply (blast dest: zadd_le_0_lemma zle_zless_trans)+
  16.919 +done
  16.920 +
  16.921 +lemma zmod_pos_neg_trivial: "[| #0 $< a;  a$+b $<= #0 |] ==> a zmod b = a$+b"
  16.922 +apply (cut_tac a = "intify (a)" and b = "intify (b)" 
  16.923 +       in zmod_pos_neg_trivial_raw)
  16.924 +apply auto
  16.925 +done
  16.926 +
  16.927 +(*There is no zmod_neg_pos_trivial...*)
  16.928 +
  16.929 +
  16.930 +(*Simpler laws such as -a zdiv b = -(a zdiv b) FAIL*)
  16.931 +
  16.932 +lemma zdiv_zminus_zminus_raw:
  16.933 +     "[|a \<in> int;  b \<in> int|] ==> ($-a) zdiv ($-b) = a zdiv b"
  16.934 +apply (case_tac "b = #0")
  16.935 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
  16.936 +apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_div])
  16.937 +apply auto
  16.938 +done
  16.939 +
  16.940 +lemma zdiv_zminus_zminus [simp]: "($-a) zdiv ($-b) = a zdiv b"
  16.941 +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zminus_zminus_raw)
  16.942 +apply auto
  16.943 +done
  16.944 +
  16.945 +(*Simpler laws such as -a zmod b = -(a zmod b) FAIL*)
  16.946 +lemma zmod_zminus_zminus_raw:
  16.947 +     "[|a \<in> int;  b \<in> int|] ==> ($-a) zmod ($-b) = $- (a zmod b)"
  16.948 +apply (case_tac "b = #0")
  16.949 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
  16.950 +apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod])
  16.951 +apply auto
  16.952 +done
  16.953 +
  16.954 +lemma zmod_zminus_zminus [simp]: "($-a) zmod ($-b) = $- (a zmod b)"
  16.955 +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zminus_zminus_raw)
  16.956 +apply auto
  16.957 +done
  16.958 +
  16.959 +
  16.960 +subsection{* division of a number by itself *}
  16.961 +
  16.962 +lemma self_quotient_aux1: "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $<= q"
  16.963 +apply (subgoal_tac "#0 $< a$*q")
  16.964 +apply (cut_tac w = "#0" and z = "q" in add1_zle_iff)
  16.965 +apply (simp add: int_0_less_mult_iff)
  16.966 +apply (blast dest: zless_trans)
  16.967 +(*linear arithmetic...*)
  16.968 +apply (drule_tac t = "%x. x $- r" in subst_context)
  16.969 +apply (drule sym)
  16.970 +apply (simp add: zcompare_rls)
  16.971 +done
  16.972 +
  16.973 +lemma self_quotient_aux2: "[| #0 $< a; a = r $+ a$*q; #0 $<= r |] ==> q $<= #1"
  16.974 +apply (subgoal_tac "#0 $<= a$* (#1$-q)")
  16.975 + apply (simp add: int_0_le_mult_iff zcompare_rls)
  16.976 + apply (blast dest: zle_zless_trans)
  16.977 +apply (simp add: zdiff_zmult_distrib2)
  16.978 +apply (drule_tac t = "%x. x $- a $* q" in subst_context)
  16.979 +apply (simp add: zcompare_rls)
  16.980 +done
  16.981 +
  16.982 +lemma self_quotient:
  16.983 +     "[| quorem(<a,a>,<q,r>);  a \<in> int;  q \<in> int;  a \<noteq> #0|] ==> q = #1"
  16.984 +apply (simp add: split_ifs quorem_def neq_iff_zless)
  16.985 +apply (rule zle_anti_sym)
  16.986 +apply safe
  16.987 +apply auto
  16.988 +prefer 4 apply (blast dest: zless_trans)
  16.989 +apply (blast dest: zless_trans)
  16.990 +apply (rule_tac [3] a = "$-a" and r = "$-r" in self_quotient_aux1)
  16.991 +apply (rule_tac a = "$-a" and r = "$-r" in self_quotient_aux2)
  16.992 +apply (rule_tac [6] zminus_equation [THEN iffD1])
  16.993 +apply (rule_tac [2] zminus_equation [THEN iffD1])
  16.994 +apply (force intro: self_quotient_aux1 self_quotient_aux2
  16.995 +  simp add: zadd_commute zmult_zminus)+
  16.996 +done
  16.997 +
  16.998 +lemma self_remainder:
  16.999 +     "[|quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0|] ==> r = #0"
 16.1000 +apply (frule self_quotient)
 16.1001 +apply (auto simp add: quorem_def)
 16.1002 +done
 16.1003 +
 16.1004 +lemma zdiv_self_raw: "[|a \<noteq> #0; a \<in> int|] ==> a zdiv a = #1"
 16.1005 +apply (blast intro: quorem_div_mod [THEN self_quotient])
 16.1006 +done
 16.1007 +
 16.1008 +lemma zdiv_self [simp]: "intify(a) \<noteq> #0 ==> a zdiv a = #1"
 16.1009 +apply (drule zdiv_self_raw)
 16.1010 +apply auto
 16.1011 +done
 16.1012 +
 16.1013 +(*Here we have 0 zmod 0 = 0, also assumed by Knuth (who puts m zmod 0 = 0) *)
 16.1014 +lemma zmod_self_raw: "a \<in> int ==> a zmod a = #0"
 16.1015 +apply (case_tac "a = #0")
 16.1016 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1017 +apply (blast intro: quorem_div_mod [THEN self_remainder])
 16.1018 +done
 16.1019 +
 16.1020 +lemma zmod_self [simp]: "a zmod a = #0"
 16.1021 +apply (cut_tac a = "intify (a)" in zmod_self_raw)
 16.1022 +apply auto
 16.1023 +done
 16.1024 +
 16.1025 +
 16.1026 +subsection{* Computation of division and remainder *}
 16.1027 +
 16.1028 +lemma zdiv_zero [simp]: "#0 zdiv b = #0"
 16.1029 +apply (simp (no_asm) add: zdiv_def divAlg_def)
 16.1030 +done
 16.1031 +
 16.1032 +lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1"
 16.1033 +apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
 16.1034 +done
 16.1035 +
 16.1036 +lemma zmod_zero [simp]: "#0 zmod b = #0"
 16.1037 +apply (simp (no_asm) add: zmod_def divAlg_def)
 16.1038 +done
 16.1039 +
 16.1040 +lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1"
 16.1041 +apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
 16.1042 +done
 16.1043 +
 16.1044 +lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1"
 16.1045 +apply (simp (no_asm_simp) add: zmod_def divAlg_def)
 16.1046 +done
 16.1047 +
 16.1048 +(** a positive, b positive **)
 16.1049 +
 16.1050 +lemma zdiv_pos_pos: "[| #0 $< a;  #0 $<= b |]  
 16.1051 +      ==> a zdiv b = fst (posDivAlg(<intify(a), intify(b)>))"
 16.1052 +apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
 16.1053 +apply (auto simp add: zle_def)
 16.1054 +done
 16.1055 +
 16.1056 +lemma zmod_pos_pos:
 16.1057 +     "[| #0 $< a;  #0 $<= b |]  
 16.1058 +      ==> a zmod b = snd (posDivAlg(<intify(a), intify(b)>))"
 16.1059 +apply (simp (no_asm_simp) add: zmod_def divAlg_def)
 16.1060 +apply (auto simp add: zle_def)
 16.1061 +done
 16.1062 +
 16.1063 +(** a negative, b positive **)
 16.1064 +
 16.1065 +lemma zdiv_neg_pos:
 16.1066 +     "[| a $< #0;  #0 $< b |]  
 16.1067 +      ==> a zdiv b = fst (negDivAlg(<intify(a), intify(b)>))"
 16.1068 +apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
 16.1069 +apply (blast dest: zle_zless_trans)
 16.1070 +done
 16.1071 +
 16.1072 +lemma zmod_neg_pos:
 16.1073 +     "[| a $< #0;  #0 $< b |]  
 16.1074 +      ==> a zmod b = snd (negDivAlg(<intify(a), intify(b)>))"
 16.1075 +apply (simp (no_asm_simp) add: zmod_def divAlg_def)
 16.1076 +apply (blast dest: zle_zless_trans)
 16.1077 +done
 16.1078 +
 16.1079 +(** a positive, b negative **)
 16.1080 +
 16.1081 +lemma zdiv_pos_neg:
 16.1082 +     "[| #0 $< a;  b $< #0 |]  
 16.1083 +      ==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))"
 16.1084 +apply (simp (no_asm_simp) add: zdiv_def divAlg_def intify_eq_0_iff_zle)
 16.1085 +apply auto
 16.1086 +apply (blast dest: zle_zless_trans)+
 16.1087 +apply (blast dest: zless_trans)
 16.1088 +apply (blast intro: zless_imp_zle)
 16.1089 +done
 16.1090 +
 16.1091 +lemma zmod_pos_neg:
 16.1092 +     "[| #0 $< a;  b $< #0 |]  
 16.1093 +      ==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))"
 16.1094 +apply (simp (no_asm_simp) add: zmod_def divAlg_def intify_eq_0_iff_zle)
 16.1095 +apply auto
 16.1096 +apply (blast dest: zle_zless_trans)+
 16.1097 +apply (blast dest: zless_trans)
 16.1098 +apply (blast intro: zless_imp_zle)
 16.1099 +done
 16.1100 +
 16.1101 +(** a negative, b negative **)
 16.1102 +
 16.1103 +lemma zdiv_neg_neg:
 16.1104 +     "[| a $< #0;  b $<= #0 |]  
 16.1105 +      ==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))"
 16.1106 +apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
 16.1107 +apply auto
 16.1108 +apply (blast dest!: zle_zless_trans)+
 16.1109 +done
 16.1110 +
 16.1111 +lemma zmod_neg_neg:
 16.1112 +     "[| a $< #0;  b $<= #0 |]  
 16.1113 +      ==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))"
 16.1114 +apply (simp (no_asm_simp) add: zmod_def divAlg_def)
 16.1115 +apply auto
 16.1116 +apply (blast dest!: zle_zless_trans)+
 16.1117 +done
 16.1118 +
 16.1119 +declare zdiv_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp]
 16.1120 +declare zdiv_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp]
 16.1121 +declare zdiv_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp]
 16.1122 +declare zdiv_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp]
 16.1123 +declare zmod_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp]
 16.1124 +declare zmod_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp]
 16.1125 +declare zmod_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp]
 16.1126 +declare zmod_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp]
 16.1127 +declare posDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp]
 16.1128 +declare negDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp]
 16.1129 +
 16.1130 +
 16.1131 +(** Special-case simplification **)
 16.1132 +
 16.1133 +lemma zmod_1 [simp]: "a zmod #1 = #0"
 16.1134 +apply (cut_tac a = "a" and b = "#1" in pos_mod_sign)
 16.1135 +apply (cut_tac [2] a = "a" and b = "#1" in pos_mod_bound)
 16.1136 +apply auto
 16.1137 +(*arithmetic*)
 16.1138 +apply (drule add1_zle_iff [THEN iffD2])
 16.1139 +apply (rule zle_anti_sym)
 16.1140 +apply auto
 16.1141 +done
 16.1142 +
 16.1143 +lemma zdiv_1 [simp]: "a zdiv #1 = intify(a)"
 16.1144 +apply (cut_tac a = "a" and b = "#1" in zmod_zdiv_equality)
 16.1145 +apply auto
 16.1146 +done
 16.1147 +
 16.1148 +lemma zmod_minus1_right [simp]: "a zmod #-1 = #0"
 16.1149 +apply (cut_tac a = "a" and b = "#-1" in neg_mod_sign)
 16.1150 +apply (cut_tac [2] a = "a" and b = "#-1" in neg_mod_bound)
 16.1151 +apply auto
 16.1152 +(*arithmetic*)
 16.1153 +apply (drule add1_zle_iff [THEN iffD2])
 16.1154 +apply (rule zle_anti_sym)
 16.1155 +apply auto
 16.1156 +done
 16.1157 +
 16.1158 +lemma zdiv_minus1_right_raw: "a \<in> int ==> a zdiv #-1 = $-a"
 16.1159 +apply (cut_tac a = "a" and b = "#-1" in zmod_zdiv_equality)
 16.1160 +apply auto
 16.1161 +apply (rule equation_zminus [THEN iffD2])
 16.1162 +apply auto
 16.1163 +done
 16.1164 +
 16.1165 +lemma zdiv_minus1_right: "a zdiv #-1 = $-a"
 16.1166 +apply (cut_tac a = "intify (a)" in zdiv_minus1_right_raw)
 16.1167 +apply auto
 16.1168 +done
 16.1169 +declare zdiv_minus1_right [simp]
 16.1170 +
 16.1171 +
 16.1172 +subsection{* Monotonicity in the first argument (divisor) *}
 16.1173 +
 16.1174 +lemma zdiv_mono1: "[| a $<= a';  #0 $< b |] ==> a zdiv b $<= a' zdiv b"
 16.1175 +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
 16.1176 +apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality)
 16.1177 +apply (rule unique_quotient_lemma)
 16.1178 +apply (erule subst)
 16.1179 +apply (erule subst)
 16.1180 +apply (simp_all (no_asm_simp) add: pos_mod_sign pos_mod_bound)
 16.1181 +done
 16.1182 +
 16.1183 +lemma zdiv_mono1_neg: "[| a $<= a';  b $< #0 |] ==> a' zdiv b $<= a zdiv b"
 16.1184 +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
 16.1185 +apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality)
 16.1186 +apply (rule unique_quotient_lemma_neg)
 16.1187 +apply (erule subst)
 16.1188 +apply (erule subst)
 16.1189 +apply (simp_all (no_asm_simp) add: neg_mod_sign neg_mod_bound)
 16.1190 +done
 16.1191 +
 16.1192 +
 16.1193 +subsection{* Monotonicity in the second argument (dividend) *}
 16.1194 +
 16.1195 +lemma q_pos_lemma:
 16.1196 +     "[| #0 $<= b'$*q' $+ r'; r' $< b';  #0 $< b' |] ==> #0 $<= q'"
 16.1197 +apply (subgoal_tac "#0 $< b'$* (q' $+ #1)")
 16.1198 + apply (simp add: int_0_less_mult_iff)
 16.1199 + apply (blast dest: zless_trans intro: zless_add1_iff_zle [THEN iffD1])
 16.1200 +apply (simp add: zadd_zmult_distrib2)
 16.1201 +apply (erule zle_zless_trans)
 16.1202 +apply (erule zadd_zless_mono2)
 16.1203 +done
 16.1204 +
 16.1205 +lemma zdiv_mono2_lemma:
 16.1206 +     "[| b$*q $+ r = b'$*q' $+ r';  #0 $<= b'$*q' $+ r';   
 16.1207 +         r' $< b';  #0 $<= r;  #0 $< b';  b' $<= b |]   
 16.1208 +      ==> q $<= q'"
 16.1209 +apply (frule q_pos_lemma, assumption+) 
 16.1210 +apply (subgoal_tac "b$*q $< b$* (q' $+ #1)")
 16.1211 + apply (simp add: zmult_zless_cancel1)
 16.1212 + apply (force dest: zless_add1_iff_zle [THEN iffD1] zless_trans zless_zle_trans)
 16.1213 +apply (subgoal_tac "b$*q = r' $- r $+ b'$*q'")
 16.1214 + prefer 2 apply (simp add: zcompare_rls)
 16.1215 +apply (simp (no_asm_simp) add: zadd_zmult_distrib2)
 16.1216 +apply (subst zadd_commute [of "b $\<times> q'"], rule zadd_zless_mono)
 16.1217 + prefer 2 apply (blast intro: zmult_zle_mono1)
 16.1218 +apply (subgoal_tac "r' $+ #0 $< b $+ r")
 16.1219 + apply (simp add: zcompare_rls)
 16.1220 +apply (rule zadd_zless_mono)
 16.1221 + apply auto
 16.1222 +apply (blast dest: zless_zle_trans)
 16.1223 +done
 16.1224 +
 16.1225 +
 16.1226 +lemma zdiv_mono2_raw:
 16.1227 +     "[| #0 $<= a;  #0 $< b';  b' $<= b;  a \<in> int |]   
 16.1228 +      ==> a zdiv b $<= a zdiv b'"
 16.1229 +apply (subgoal_tac "#0 $< b")
 16.1230 + prefer 2 apply (blast dest: zless_zle_trans)
 16.1231 +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
 16.1232 +apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality)
 16.1233 +apply (rule zdiv_mono2_lemma)
 16.1234 +apply (erule subst)
 16.1235 +apply (erule subst)
 16.1236 +apply (simp_all add: pos_mod_sign pos_mod_bound)
 16.1237 +done
 16.1238 +
 16.1239 +lemma zdiv_mono2:
 16.1240 +     "[| #0 $<= a;  #0 $< b';  b' $<= b |]   
 16.1241 +      ==> a zdiv b $<= a zdiv b'"
 16.1242 +apply (cut_tac a = "intify (a)" in zdiv_mono2_raw)
 16.1243 +apply auto
 16.1244 +done
 16.1245 +
 16.1246 +lemma q_neg_lemma:
 16.1247 +     "[| b'$*q' $+ r' $< #0;  #0 $<= r';  #0 $< b' |] ==> q' $< #0"
 16.1248 +apply (subgoal_tac "b'$*q' $< #0")
 16.1249 + prefer 2 apply (force intro: zle_zless_trans)
 16.1250 +apply (simp add: zmult_less_0_iff)
 16.1251 +apply (blast dest: zless_trans)
 16.1252 +done
 16.1253 +
 16.1254 +
 16.1255 +
 16.1256 +lemma zdiv_mono2_neg_lemma:
 16.1257 +     "[| b$*q $+ r = b'$*q' $+ r';  b'$*q' $+ r' $< #0;   
 16.1258 +         r $< b;  #0 $<= r';  #0 $< b';  b' $<= b |]   
 16.1259 +      ==> q' $<= q"
 16.1260 +apply (subgoal_tac "#0 $< b")
 16.1261 + prefer 2 apply (blast dest: zless_zle_trans)
 16.1262 +apply (frule q_neg_lemma, assumption+) 
 16.1263 +apply (subgoal_tac "b$*q' $< b$* (q $+ #1)")
 16.1264 + apply (simp add: zmult_zless_cancel1)
 16.1265 + apply (blast dest: zless_trans zless_add1_iff_zle [THEN iffD1])
 16.1266 +apply (simp (no_asm_simp) add: zadd_zmult_distrib2)
 16.1267 +apply (subgoal_tac "b$*q' $<= b'$*q'")
 16.1268 + prefer 2
 16.1269 + apply (simp add: zmult_zle_cancel2)
 16.1270 + apply (blast dest: zless_trans)
 16.1271 +apply (subgoal_tac "b'$*q' $+ r $< b $+ (b$*q $+ r)")
 16.1272 + prefer 2
 16.1273 + apply (erule ssubst)
 16.1274 + apply simp
 16.1275 + apply (drule_tac w' = "r" and z' = "#0" in zadd_zless_mono)
 16.1276 +  apply (assumption)
 16.1277 + apply simp
 16.1278 +apply (simp (no_asm_use) add: zadd_commute)
 16.1279 +apply (rule zle_zless_trans)
 16.1280 + prefer 2 apply (assumption)
 16.1281 +apply (simp (no_asm_simp) add: zmult_zle_cancel2)
 16.1282 +apply (blast dest: zless_trans)
 16.1283 +done
 16.1284 +
 16.1285 +lemma zdiv_mono2_neg_raw:
 16.1286 +     "[| a $< #0;  #0 $< b';  b' $<= b;  a \<in> int |]   
 16.1287 +      ==> a zdiv b' $<= a zdiv b"
 16.1288 +apply (subgoal_tac "#0 $< b")
 16.1289 + prefer 2 apply (blast dest: zless_zle_trans)
 16.1290 +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
 16.1291 +apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality)
 16.1292 +apply (rule zdiv_mono2_neg_lemma)
 16.1293 +apply (erule subst)
 16.1294 +apply (erule subst)
 16.1295 +apply (simp_all add: pos_mod_sign pos_mod_bound)
 16.1296 +done
 16.1297 +
 16.1298 +lemma zdiv_mono2_neg: "[| a $< #0;  #0 $< b';  b' $<= b |]   
 16.1299 +      ==> a zdiv b' $<= a zdiv b"
 16.1300 +apply (cut_tac a = "intify (a)" in zdiv_mono2_neg_raw)
 16.1301 +apply auto
 16.1302 +done
 16.1303 +
 16.1304 +
 16.1305 +
 16.1306 +subsection{* More algebraic laws for zdiv and zmod *}
 16.1307 +
 16.1308 +(** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **)
 16.1309 +
 16.1310 +lemma zmult1_lemma:
 16.1311 +     "[| quorem(<b,c>, <q,r>);  c \<in> int;  c \<noteq> #0 |]  
 16.1312 +      ==> quorem (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)"
 16.1313 +apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2
 16.1314 +                      pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound)
 16.1315 +apply (auto intro: raw_zmod_zdiv_equality) 
 16.1316 +done
 16.1317 +
 16.1318 +lemma zdiv_zmult1_eq_raw:
 16.1319 +     "[|b \<in> int;  c \<in> int|]  
 16.1320 +      ==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c"
 16.1321 +apply (case_tac "c = #0")
 16.1322 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1323 +apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_div])
 16.1324 +apply auto
 16.1325 +done
 16.1326 +
 16.1327 +lemma zdiv_zmult1_eq: "(a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c"
 16.1328 +apply (cut_tac b = "intify (b)" and c = "intify (c)" in zdiv_zmult1_eq_raw)
 16.1329 +apply auto
 16.1330 +done
 16.1331 +
 16.1332 +lemma zmod_zmult1_eq_raw:
 16.1333 +     "[|b \<in> int;  c \<in> int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c"
 16.1334 +apply (case_tac "c = #0")
 16.1335 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1336 +apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod])
 16.1337 +apply auto
 16.1338 +done
 16.1339 +
 16.1340 +lemma zmod_zmult1_eq: "(a$*b) zmod c = a$*(b zmod c) zmod c"
 16.1341 +apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult1_eq_raw)
 16.1342 +apply auto
 16.1343 +done
 16.1344 +
 16.1345 +lemma zmod_zmult1_eq': "(a$*b) zmod c = ((a zmod c) $* b) zmod c"
 16.1346 +apply (rule trans)
 16.1347 +apply (rule_tac b = " (b $* a) zmod c" in trans)
 16.1348 +apply (rule_tac [2] zmod_zmult1_eq)
 16.1349 +apply (simp_all (no_asm) add: zmult_commute)
 16.1350 +done
 16.1351 +
 16.1352 +lemma zmod_zmult_distrib: "(a$*b) zmod c = ((a zmod c) $* (b zmod c)) zmod c"
 16.1353 +apply (rule zmod_zmult1_eq' [THEN trans])
 16.1354 +apply (rule zmod_zmult1_eq)
 16.1355 +done
 16.1356 +
 16.1357 +lemma zdiv_zmult_self1 [simp]: "intify(b) \<noteq> #0 ==> (a$*b) zdiv b = intify(a)"
 16.1358 +apply (simp (no_asm_simp) add: zdiv_zmult1_eq)
 16.1359 +done
 16.1360 +
 16.1361 +lemma zdiv_zmult_self2 [simp]: "intify(b) \<noteq> #0 ==> (b$*a) zdiv b = intify(a)"
 16.1362 +apply (subst zmult_commute , erule zdiv_zmult_self1)
 16.1363 +done
 16.1364 +
 16.1365 +lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0"
 16.1366 +apply (simp (no_asm) add: zmod_zmult1_eq)
 16.1367 +done
 16.1368 +
 16.1369 +lemma zmod_zmult_self2 [simp]: "(b$*a) zmod b = #0"
 16.1370 +apply (simp (no_asm) add: zmult_commute zmod_zmult1_eq)
 16.1371 +done
 16.1372 +
 16.1373 +
 16.1374 +(** proving (a$+b) zdiv c = 
 16.1375 +            a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **)
 16.1376 +
 16.1377 +lemma zadd1_lemma:
 16.1378 +     "[| quorem(<a,c>, <aq,ar>);  quorem(<b,c>, <bq,br>);   
 16.1379 +         c \<in> int;  c \<noteq> #0 |]  
 16.1380 +      ==> quorem (<a$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)"
 16.1381 +apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2
 16.1382 +                      pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound)
 16.1383 +apply (auto intro: raw_zmod_zdiv_equality)
 16.1384 +done
 16.1385 +
 16.1386 +(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
 16.1387 +lemma zdiv_zadd1_eq_raw:
 16.1388 +     "[|a \<in> int; b \<in> int; c \<in> int|] ==>  
 16.1389 +      (a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)"
 16.1390 +apply (case_tac "c = #0")
 16.1391 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1392 +apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod,
 16.1393 +                                 THEN quorem_div])
 16.1394 +done
 16.1395 +
 16.1396 +lemma zdiv_zadd1_eq:
 16.1397 +     "(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)"
 16.1398 +apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" 
 16.1399 +       in zdiv_zadd1_eq_raw)
 16.1400 +apply auto
 16.1401 +done
 16.1402 +
 16.1403 +lemma zmod_zadd1_eq_raw:
 16.1404 +     "[|a \<in> int; b \<in> int; c \<in> int|]   
 16.1405 +      ==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c"
 16.1406 +apply (case_tac "c = #0")
 16.1407 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1408 +apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, 
 16.1409 +                                 THEN quorem_mod])
 16.1410 +done
 16.1411 +
 16.1412 +lemma zmod_zadd1_eq: "(a$+b) zmod c = (a zmod c $+ b zmod c) zmod c"
 16.1413 +apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" 
 16.1414 +       in zmod_zadd1_eq_raw)
 16.1415 +apply auto
 16.1416 +done
 16.1417 +
 16.1418 +lemma zmod_div_trivial_raw:
 16.1419 +     "[|a \<in> int; b \<in> int|] ==> (a zmod b) zdiv b = #0"
 16.1420 +apply (case_tac "b = #0")
 16.1421 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1422 +apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound
 16.1423 +         zdiv_pos_pos_trivial neg_mod_sign neg_mod_bound zdiv_neg_neg_trivial)
 16.1424 +done
 16.1425 +
 16.1426 +lemma zmod_div_trivial [simp]: "(a zmod b) zdiv b = #0"
 16.1427 +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_div_trivial_raw)
 16.1428 +apply auto
 16.1429 +done
 16.1430 +
 16.1431 +lemma zmod_mod_trivial_raw:
 16.1432 +     "[|a \<in> int; b \<in> int|] ==> (a zmod b) zmod b = a zmod b"
 16.1433 +apply (case_tac "b = #0")
 16.1434 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1435 +apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound 
 16.1436 +       zmod_pos_pos_trivial neg_mod_sign neg_mod_bound zmod_neg_neg_trivial)
 16.1437 +done
 16.1438 +
 16.1439 +lemma zmod_mod_trivial [simp]: "(a zmod b) zmod b = a zmod b"
 16.1440 +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_mod_trivial_raw)
 16.1441 +apply auto
 16.1442 +done
 16.1443 +
 16.1444 +lemma zmod_zadd_left_eq: "(a$+b) zmod c = ((a zmod c) $+ b) zmod c"
 16.1445 +apply (rule trans [symmetric])
 16.1446 +apply (rule zmod_zadd1_eq)
 16.1447 +apply (simp (no_asm))
 16.1448 +apply (rule zmod_zadd1_eq [symmetric])
 16.1449 +done
 16.1450 +
 16.1451 +lemma zmod_zadd_right_eq: "(a$+b) zmod c = (a $+ (b zmod c)) zmod c"
 16.1452 +apply (rule trans [symmetric])
 16.1453 +apply (rule zmod_zadd1_eq)
 16.1454 +apply (simp (no_asm))
 16.1455 +apply (rule zmod_zadd1_eq [symmetric])
 16.1456 +done
 16.1457 +
 16.1458 +
 16.1459 +lemma zdiv_zadd_self1 [simp]:
 16.1460 +     "intify(a) \<noteq> #0 ==> (a$+b) zdiv a = b zdiv a $+ #1"
 16.1461 +by (simp (no_asm_simp) add: zdiv_zadd1_eq)
 16.1462 +
 16.1463 +lemma zdiv_zadd_self2 [simp]:
 16.1464 +     "intify(a) \<noteq> #0 ==> (b$+a) zdiv a = b zdiv a $+ #1"
 16.1465 +by (simp (no_asm_simp) add: zdiv_zadd1_eq)
 16.1466 +
 16.1467 +lemma zmod_zadd_self1 [simp]: "(a$+b) zmod a = b zmod a"
 16.1468 +apply (case_tac "a = #0")
 16.1469 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1470 +apply (simp (no_asm_simp) add: zmod_zadd1_eq)
 16.1471 +done
 16.1472 +
 16.1473 +lemma zmod_zadd_self2 [simp]: "(b$+a) zmod a = b zmod a"
 16.1474 +apply (case_tac "a = #0")
 16.1475 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1476 +apply (simp (no_asm_simp) add: zmod_zadd1_eq)
 16.1477 +done
 16.1478 +
 16.1479 +
 16.1480 +subsection{* proving  a zdiv (b*c) = (a zdiv b) zdiv c *}
 16.1481 +
 16.1482 +(*The condition c>0 seems necessary.  Consider that 7 zdiv ~6 = ~2 but
 16.1483 +  7 zdiv 2 zdiv ~3 = 3 zdiv ~3 = ~1.  The subcase (a zdiv b) zmod c = 0 seems
 16.1484 +  to cause particular problems.*)
 16.1485 +
 16.1486 +(** first, four lemmas to bound the remainder for the cases b<0 and b>0 **)
 16.1487 +
 16.1488 +lemma zdiv_zmult2_aux1:
 16.1489 +     "[| #0 $< c;  b $< r;  r $<= #0 |] ==> b$*c $< b$*(q zmod c) $+ r"
 16.1490 +apply (subgoal_tac "b $* (c $- q zmod c) $< r $* #1")
 16.1491 +apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls)
 16.1492 +apply (rule zle_zless_trans)
 16.1493 +apply (erule_tac [2] zmult_zless_mono1)
 16.1494 +apply (rule zmult_zle_mono2_neg)
 16.1495 +apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound)
 16.1496 +apply (blast intro: zless_imp_zle dest: zless_zle_trans)
 16.1497 +done
 16.1498 +
 16.1499 +lemma zdiv_zmult2_aux2:
 16.1500 +     "[| #0 $< c;   b $< r;  r $<= #0 |] ==> b $* (q zmod c) $+ r $<= #0"
 16.1501 +apply (subgoal_tac "b $* (q zmod c) $<= #0")
 16.1502 + prefer 2
 16.1503 + apply (simp add: zmult_le_0_iff pos_mod_sign) 
 16.1504 + apply (blast intro: zless_imp_zle dest: zless_zle_trans)
 16.1505 +(*arithmetic*)
 16.1506 +apply (drule zadd_zle_mono)
 16.1507 +apply assumption
 16.1508 +apply (simp add: zadd_commute)
 16.1509 +done
 16.1510 +
 16.1511 +lemma zdiv_zmult2_aux3:
 16.1512 +     "[| #0 $< c;  #0 $<= r;  r $< b |] ==> #0 $<= b $* (q zmod c) $+ r"
 16.1513 +apply (subgoal_tac "#0 $<= b $* (q zmod c)")
 16.1514 + prefer 2
 16.1515 + apply (simp add: int_0_le_mult_iff pos_mod_sign) 
 16.1516 + apply (blast intro: zless_imp_zle dest: zle_zless_trans)
 16.1517 +(*arithmetic*)
 16.1518 +apply (drule zadd_zle_mono)
 16.1519 +apply assumption
 16.1520 +apply (simp add: zadd_commute)
 16.1521 +done
 16.1522 +
 16.1523 +lemma zdiv_zmult2_aux4:
 16.1524 +     "[| #0 $< c; #0 $<= r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c"
 16.1525 +apply (subgoal_tac "r $* #1 $< b $* (c $- q zmod c)")
 16.1526 +apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls)
 16.1527 +apply (rule zless_zle_trans)
 16.1528 +apply (erule zmult_zless_mono1)
 16.1529 +apply (rule_tac [2] zmult_zle_mono2)
 16.1530 +apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound)
 16.1531 +apply (blast intro: zless_imp_zle dest: zle_zless_trans)
 16.1532 +done
 16.1533 +
 16.1534 +lemma zdiv_zmult2_lemma:
 16.1535 +     "[| quorem (<a,b>, <q,r>);  a \<in> int;  b \<in> int;  b \<noteq> #0;  #0 $< c |]  
 16.1536 +      ==> quorem (<a,b$*c>, <q zdiv c, b$*(q zmod c) $+ r>)"
 16.1537 +apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def
 16.1538 +               neq_iff_zless int_0_less_mult_iff 
 16.1539 +               zadd_zmult_distrib2 [symmetric] zdiv_zmult2_aux1 zdiv_zmult2_aux2
 16.1540 +               zdiv_zmult2_aux3 zdiv_zmult2_aux4)
 16.1541 +apply (blast dest: zless_trans)+
 16.1542 +done
 16.1543 +
 16.1544 +lemma zdiv_zmult2_eq_raw:
 16.1545 +     "[|#0 $< c;  a \<in> int;  b \<in> int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c"
 16.1546 +apply (case_tac "b = #0")
 16.1547 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1548 +apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_div])
 16.1549 +apply (auto simp add: intify_eq_0_iff_zle)
 16.1550 +apply (blast dest: zle_zless_trans)
 16.1551 +done
 16.1552 +
 16.1553 +lemma zdiv_zmult2_eq: "#0 $< c ==> a zdiv (b$*c) = (a zdiv b) zdiv c"
 16.1554 +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zmult2_eq_raw)
 16.1555 +apply auto
 16.1556 +done
 16.1557 +
 16.1558 +lemma zmod_zmult2_eq_raw:
 16.1559 +     "[|#0 $< c;  a \<in> int;  b \<in> int|]  
 16.1560 +      ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b"
 16.1561 +apply (case_tac "b = #0")
 16.1562 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1563 +apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_mod])
 16.1564 +apply (auto simp add: intify_eq_0_iff_zle)
 16.1565 +apply (blast dest: zle_zless_trans)
 16.1566 +done
 16.1567 +
 16.1568 +lemma zmod_zmult2_eq:
 16.1569 +     "#0 $< c ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b"
 16.1570 +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zmult2_eq_raw)
 16.1571 +apply auto
 16.1572 +done
 16.1573 +
 16.1574 +subsection{* Cancellation of common factors in "zdiv" *}
 16.1575 +
 16.1576 +lemma zdiv_zmult_zmult1_aux1:
 16.1577 +     "[| #0 $< b;  intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b"
 16.1578 +apply (subst zdiv_zmult2_eq)
 16.1579 +apply auto
 16.1580 +done
 16.1581 +
 16.1582 +lemma zdiv_zmult_zmult1_aux2:
 16.1583 +     "[| b $< #0;  intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b"
 16.1584 +apply (subgoal_tac " (c $* ($-a)) zdiv (c $* ($-b)) = ($-a) zdiv ($-b)")
 16.1585 +apply (rule_tac [2] zdiv_zmult_zmult1_aux1)
 16.1586 +apply auto
 16.1587 +done
 16.1588 +
 16.1589 +lemma zdiv_zmult_zmult1_raw:
 16.1590 +     "[|intify(c) \<noteq> #0; b \<in> int|] ==> (c$*a) zdiv (c$*b) = a zdiv b"
 16.1591 +apply (case_tac "b = #0")
 16.1592 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1593 +apply (auto simp add: neq_iff_zless [of b]
 16.1594 +  zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2)
 16.1595 +done
 16.1596 +
 16.1597 +lemma zdiv_zmult_zmult1: "intify(c) \<noteq> #0 ==> (c$*a) zdiv (c$*b) = a zdiv b"
 16.1598 +apply (cut_tac b = "intify (b)" in zdiv_zmult_zmult1_raw)
 16.1599 +apply auto
 16.1600 +done
 16.1601 +
 16.1602 +lemma zdiv_zmult_zmult2: "intify(c) \<noteq> #0 ==> (a$*c) zdiv (b$*c) = a zdiv b"
 16.1603 +apply (drule zdiv_zmult_zmult1)
 16.1604 +apply (auto simp add: zmult_commute)
 16.1605 +done
 16.1606 +
 16.1607 +
 16.1608 +subsection{* Distribution of factors over "zmod" *}
 16.1609 +
 16.1610 +lemma zmod_zmult_zmult1_aux1:
 16.1611 +     "[| #0 $< b;  intify(c) \<noteq> #0 |]  
 16.1612 +      ==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
 16.1613 +apply (subst zmod_zmult2_eq)
 16.1614 +apply auto
 16.1615 +done
 16.1616 +
 16.1617 +lemma zmod_zmult_zmult1_aux2:
 16.1618 +     "[| b $< #0;  intify(c) \<noteq> #0 |]  
 16.1619 +      ==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
 16.1620 +apply (subgoal_tac " (c $* ($-a)) zmod (c $* ($-b)) = c $* (($-a) zmod ($-b))")
 16.1621 +apply (rule_tac [2] zmod_zmult_zmult1_aux1)
 16.1622 +apply auto
 16.1623 +done
 16.1624 +
 16.1625 +lemma zmod_zmult_zmult1_raw:
 16.1626 +     "[|b \<in> int; c \<in> int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
 16.1627 +apply (case_tac "b = #0")
 16.1628 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1629 +apply (case_tac "c = #0")
 16.1630 + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) 
 16.1631 +apply (auto simp add: neq_iff_zless [of b]
 16.1632 +  zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2)
 16.1633 +done
 16.1634 +
 16.1635 +lemma zmod_zmult_zmult1: "(c$*a) zmod (c$*b) = c $* (a zmod b)"
 16.1636 +apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult_zmult1_raw)
 16.1637 +apply auto
 16.1638 +done
 16.1639 +
 16.1640 +lemma zmod_zmult_zmult2: "(a$*c) zmod (b$*c) = (a zmod b) $* c"
 16.1641 +apply (cut_tac c = "c" in zmod_zmult_zmult1)
 16.1642 +apply (auto simp add: zmult_commute)
 16.1643 +done
 16.1644 +
 16.1645 +
 16.1646 +(** Quotients of signs **)
 16.1647 +
 16.1648 +lemma zdiv_neg_pos_less0: "[| a $< #0;  #0 $< b |] ==> a zdiv b $< #0"
 16.1649 +apply (subgoal_tac "a zdiv b $<= #-1")
 16.1650 +apply (erule zle_zless_trans)
 16.1651 +apply (simp (no_asm))
 16.1652 +apply (rule zle_trans)
 16.1653 +apply (rule_tac a' = "#-1" in zdiv_mono1)
 16.1654 +apply (rule zless_add1_iff_zle [THEN iffD1])
 16.1655 +apply (simp (no_asm))
 16.1656 +apply (auto simp add: zdiv_minus1)
 16.1657 +done
 16.1658 +
 16.1659 +lemma zdiv_nonneg_neg_le0: "[| #0 $<= a;  b $< #0 |] ==> a zdiv b $<= #0"
 16.1660 +apply (drule zdiv_mono1_neg)
 16.1661 +apply auto
 16.1662 +done
 16.1663 +
 16.1664 +lemma pos_imp_zdiv_nonneg_iff: "#0 $< b ==> (#0 $<= a zdiv b) <-> (#0 $<= a)"
 16.1665 +apply auto
 16.1666 +apply (drule_tac [2] zdiv_mono1)
 16.1667 +apply (auto simp add: neq_iff_zless)
 16.1668 +apply (simp (no_asm_use) add: not_zless_iff_zle [THEN iff_sym])
 16.1669 +apply (blast intro: zdiv_neg_pos_less0)
 16.1670 +done
 16.1671 +
 16.1672 +lemma neg_imp_zdiv_nonneg_iff: "b $< #0 ==> (#0 $<= a zdiv b) <-> (a $<= #0)"
 16.1673 +apply (subst zdiv_zminus_zminus [symmetric])
 16.1674 +apply (rule iff_trans)
 16.1675 +apply (rule pos_imp_zdiv_nonneg_iff)
 16.1676 +apply auto
 16.1677 +done
 16.1678 +
 16.1679 +(*But not (a zdiv b $<= 0 iff a$<=0); consider a=1, b=2 when a zdiv b = 0.*)
 16.1680 +lemma pos_imp_zdiv_neg_iff: "#0 $< b ==> (a zdiv b $< #0) <-> (a $< #0)"
 16.1681 +apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym])
 16.1682 +apply (erule pos_imp_zdiv_nonneg_iff)
 16.1683 +done
 16.1684 +
 16.1685 +(*Again the law fails for $<=: consider a = -1, b = -2 when a zdiv b = 0*)
 16.1686 +lemma neg_imp_zdiv_neg_iff: "b $< #0 ==> (a zdiv b $< #0) <-> (#0 $< a)"
 16.1687 +apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym])
 16.1688 +apply (erule neg_imp_zdiv_nonneg_iff)
 16.1689 +done
 16.1690 +
 16.1691 +(*
 16.1692 + THESE REMAIN TO BE CONVERTED -- but aren't that useful!
 16.1693 +
 16.1694 + subsection{* Speeding up the division algorithm with shifting *}
 16.1695 +
 16.1696 + (** computing "zdiv" by shifting **)
 16.1697 +
 16.1698 + lemma pos_zdiv_mult_2: "#0 $<= a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a"
 16.1699 + apply (case_tac "a = #0")
 16.1700 + apply (subgoal_tac "#1 $<= a")
 16.1701 +  apply (arith_tac 2)
 16.1702 + apply (subgoal_tac "#1 $< a $* #2")
 16.1703 +  apply (arith_tac 2)
 16.1704 + apply (subgoal_tac "#2$* (#1 $+ b zmod a) $<= #2$*a")
 16.1705 +  apply (rule_tac [2] zmult_zle_mono2)
 16.1706 + apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound)
 16.1707 + apply (subst zdiv_zadd1_eq)
 16.1708 + apply (simp (no_asm_simp) add: zdiv_zmult_zmult2 zmod_zmult_zmult2 zdiv_pos_pos_trivial)
 16.1709 + apply (subst zdiv_pos_pos_trivial)
 16.1710 + apply (simp (no_asm_simp) add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ])
 16.1711 + apply (auto simp add: zmod_pos_pos_trivial)
 16.1712 + apply (subgoal_tac "#0 $<= b zmod a")
 16.1713 +  apply (asm_simp_tac (simpset () add: pos_mod_sign) 2)
 16.1714 + apply arith
 16.1715 + done
 16.1716 +
 16.1717 +
 16.1718 + lemma neg_zdiv_mult_2: "a $<= #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) <-> (b$+#1) zdiv a"
 16.1719 + apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zdiv (#2 $* ($-a)) <-> ($-b-#1) zdiv ($-a)")
 16.1720 + apply (rule_tac [2] pos_zdiv_mult_2)
 16.1721 + apply (auto simp add: zmult_zminus_right)
 16.1722 + apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))")
 16.1723 + apply (Simp_tac 2)
 16.1724 + apply (asm_full_simp_tac (HOL_ss add: zdiv_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric])
 16.1725 + done
 16.1726 +
 16.1727 +
 16.1728 + (*Not clear why this must be proved separately; probably integ_of causes
 16.1729 +   simplification problems*)
 16.1730 + lemma lemma: "~ #0 $<= x ==> x $<= #0"
 16.1731 + apply auto
 16.1732 + done
 16.1733 +
 16.1734 + lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) =  
 16.1735 +           (if ~b | #0 $<= integ_of w                    
 16.1736 +            then integ_of v zdiv (integ_of w)     
 16.1737 +            else (integ_of v $+ #1) zdiv (integ_of w))"
 16.1738 + apply (simp_tac (simpset_of Int.thy add: zadd_assoc integ_of_BIT)
 16.1739 + 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)
 16.1740 + done
 16.1741 +
 16.1742 + declare zdiv_integ_of_BIT [simp]
 16.1743 +
 16.1744 +
 16.1745 + (** computing "zmod" by shifting (proofs resemble those for "zdiv") **)
 16.1746 +
 16.1747 + lemma pos_zmod_mult_2: "#0 $<= a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)"
 16.1748 + apply (case_tac "a = #0")
 16.1749 + apply (subgoal_tac "#1 $<= a")
 16.1750 +  apply (arith_tac 2)
 16.1751 + apply (subgoal_tac "#1 $< a $* #2")
 16.1752 +  apply (arith_tac 2)
 16.1753 + apply (subgoal_tac "#2$* (#1 $+ b zmod a) $<= #2$*a")
 16.1754 +  apply (rule_tac [2] zmult_zle_mono2)
 16.1755 + apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound)
 16.1756 + apply (subst zmod_zadd1_eq)
 16.1757 + apply (simp (no_asm_simp) add: zmod_zmult_zmult2 zmod_pos_pos_trivial)
 16.1758 + apply (rule zmod_pos_pos_trivial)
 16.1759 + apply (simp (no_asm_simp) # add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ])
 16.1760 + apply (auto simp add: zmod_pos_pos_trivial)
 16.1761 + apply (subgoal_tac "#0 $<= b zmod a")
 16.1762 +  apply (asm_simp_tac (simpset () add: pos_mod_sign) 2)
 16.1763 + apply arith
 16.1764 + done
 16.1765 +
 16.1766 +
 16.1767 + lemma neg_zmod_mult_2: "a $<= #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1"
 16.1768 + apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zmod (#2$* ($-a)) = #1 $+ #2$* (($-b-#1) zmod ($-a))")
 16.1769 + apply (rule_tac [2] pos_zmod_mult_2)
 16.1770 + apply (auto simp add: zmult_zminus_right)
 16.1771 + apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))")
 16.1772 + apply (Simp_tac 2)
 16.1773 + apply (asm_full_simp_tac (HOL_ss add: zmod_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric])
 16.1774 + apply (dtac (zminus_equation [THEN iffD1, symmetric])
 16.1775 + apply auto
 16.1776 + done
 16.1777 +
 16.1778 + lemma zmod_integ_of_BIT: "integ_of (v BIT b) zmod integ_of (w BIT False) =  
 16.1779 +           (if b then  
 16.1780 +                 if #0 $<= integ_of w  
 16.1781 +                 then #2 $* (integ_of v zmod integ_of w) $+ #1     
 16.1782 +                 else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1   
 16.1783 +            else #2 $* (integ_of v zmod integ_of w))"
 16.1784 + apply (simp_tac (simpset_of Int.thy add: zadd_assoc integ_of_BIT)
 16.1785 + 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)
 16.1786 + done
 16.1787 +
 16.1788 + declare zmod_integ_of_BIT [simp]
 16.1789 +*)
 16.1790 +
 16.1791 +end
 16.1792 +
    17.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    17.2 +++ b/src/ZF/Int_ZF.thy	Mon Feb 11 15:40:21 2008 +0100
    17.3 @@ -0,0 +1,931 @@
    17.4 +(*  Title:      ZF/Int.thy
    17.5 +    ID:         $Id$
    17.6 +    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    17.7 +    Copyright   1993  University of Cambridge
    17.8 +
    17.9 +*)
   17.10 +
   17.11 +header{*The Integers as Equivalence Classes Over Pairs of Natural Numbers*}
   17.12 +
   17.13 +theory Int_ZF imports EquivClass ArithSimp begin
   17.14 +
   17.15 +definition
   17.16 +  intrel :: i  where
   17.17 +    "intrel == {p : (nat*nat)*(nat*nat).                 
   17.18 +                \<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1}"
   17.19 +
   17.20 +definition
   17.21 +  int :: i  where
   17.22 +    "int == (nat*nat)//intrel"  
   17.23 +
   17.24 +definition
   17.25 +  int_of :: "i=>i" --{*coercion from nat to int*}    ("$# _" [80] 80)  where
   17.26 +    "$# m == intrel `` {<natify(m), 0>}"
   17.27 +
   17.28 +definition
   17.29 +  intify :: "i=>i" --{*coercion from ANYTHING to int*}  where
   17.30 +    "intify(m) == if m : int then m else $#0"
   17.31 +
   17.32 +definition
   17.33 +  raw_zminus :: "i=>i"  where
   17.34 +    "raw_zminus(z) == \<Union><x,y>\<in>z. intrel``{<y,x>}"
   17.35 +
   17.36 +definition
   17.37 +  zminus :: "i=>i"                                 ("$- _" [80] 80)  where
   17.38 +    "$- z == raw_zminus (intify(z))"
   17.39 +
   17.40 +definition
   17.41 +  znegative   ::      "i=>o"  where
   17.42 +    "znegative(z) == \<exists>x y. x<y & y\<in>nat & <x,y>\<in>z"
   17.43 +
   17.44 +definition
   17.45 +  iszero      ::      "i=>o"  where
   17.46 +    "iszero(z) == z = $# 0"
   17.47 +    
   17.48 +definition
   17.49 +  raw_nat_of  :: "i=>i"  where
   17.50 +  "raw_nat_of(z) == natify (\<Union><x,y>\<in>z. x#-y)"
   17.51 +
   17.52 +definition
   17.53 +  nat_of  :: "i=>i"  where
   17.54 +  "nat_of(z) == raw_nat_of (intify(z))"
   17.55 +
   17.56 +definition
   17.57 +  zmagnitude  ::      "i=>i"  where
   17.58 +  --{*could be replaced by an absolute value function from int to int?*}
   17.59 +    "zmagnitude(z) ==
   17.60 +     THE m. m\<in>nat & ((~ znegative(z) & z = $# m) |
   17.61 +		       (znegative(z) & $- z = $# m))"
   17.62 +
   17.63 +definition
   17.64 +  raw_zmult   ::      "[i,i]=>i"  where
   17.65 +    (*Cannot use UN<x1,y2> here or in zadd because of the form of congruent2.
   17.66 +      Perhaps a "curried" or even polymorphic congruent predicate would be
   17.67 +      better.*)
   17.68 +     "raw_zmult(z1,z2) == 
   17.69 +       \<Union>p1\<in>z1. \<Union>p2\<in>z2.  split(%x1 y1. split(%x2 y2.        
   17.70 +                   intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)"
   17.71 +
   17.72 +definition
   17.73 +  zmult       ::      "[i,i]=>i"      (infixl "$*" 70)  where
   17.74 +     "z1 $* z2 == raw_zmult (intify(z1),intify(z2))"
   17.75 +
   17.76 +definition
   17.77 +  raw_zadd    ::      "[i,i]=>i"  where
   17.78 +     "raw_zadd (z1, z2) == 
   17.79 +       \<Union>z1\<in>z1. \<Union>z2\<in>z2. let <x1,y1>=z1; <x2,y2>=z2                 
   17.80 +                           in intrel``{<x1#+x2, y1#+y2>}"
   17.81 +
   17.82 +definition
   17.83 +  zadd        ::      "[i,i]=>i"      (infixl "$+" 65)  where
   17.84 +     "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))"
   17.85 +
   17.86 +definition
   17.87 +  zdiff        ::      "[i,i]=>i"      (infixl "$-" 65)  where
   17.88 +     "z1 $- z2 == z1 $+ zminus(z2)"
   17.89 +
   17.90 +definition
   17.91 +  zless        ::      "[i,i]=>o"      (infixl "$<" 50)  where
   17.92 +     "z1 $< z2 == znegative(z1 $- z2)"
   17.93 +  
   17.94 +definition
   17.95 +  zle          ::      "[i,i]=>o"      (infixl "$<=" 50)  where
   17.96 +     "z1 $<= z2 == z1 $< z2 | intify(z1)=intify(z2)"
   17.97 +  
   17.98 +
   17.99 +notation (xsymbols)
  17.100 +  zmult  (infixl "$\<times>" 70) and
  17.101 +  zle  (infixl "$\<le>" 50)  --{*less than or equals*}
  17.102 +
  17.103 +notation (HTML output)
  17.104 +  zmult  (infixl "$\<times>" 70) and
  17.105 +  zle  (infixl "$\<le>" 50)
  17.106 +
  17.107 +
  17.108 +declare quotientE [elim!]
  17.109 +
  17.110 +subsection{*Proving that @{term intrel} is an equivalence relation*}
  17.111 +
  17.112 +(** Natural deduction for intrel **)
  17.113 +
  17.114 +lemma intrel_iff [simp]: 
  17.115 +    "<<x1,y1>,<x2,y2>>: intrel <->  
  17.116 +     x1\<in>nat & y1\<in>nat & x2\<in>nat & y2\<in>nat & x1#+y2 = x2#+y1"
  17.117 +by (simp add: intrel_def)
  17.118 +
  17.119 +lemma intrelI [intro!]: 
  17.120 +    "[| x1#+y2 = x2#+y1; x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |]   
  17.121 +     ==> <<x1,y1>,<x2,y2>>: intrel"
  17.122 +by (simp add: intrel_def)
  17.123 +
  17.124 +lemma intrelE [elim!]:
  17.125 +  "[| p: intrel;   
  17.126 +      !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>;  x1#+y2 = x2#+y1;  
  17.127 +                        x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] ==> Q |]  
  17.128 +   ==> Q"
  17.129 +by (simp add: intrel_def, blast) 
  17.130 +
  17.131 +lemma int_trans_lemma:
  17.132 +     "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2 |] ==> x1 #+ y3 = x3 #+ y1"
  17.133 +apply (rule sym)
  17.134 +apply (erule add_left_cancel)+
  17.135 +apply (simp_all (no_asm_simp))
  17.136 +done
  17.137 +
  17.138 +lemma equiv_intrel: "equiv(nat*nat, intrel)"
  17.139 +apply (simp add: equiv_def refl_def sym_def trans_def)
  17.140 +apply (fast elim!: sym int_trans_lemma)
  17.141 +done
  17.142 +
  17.143 +lemma image_intrel_int: "[| m\<in>nat; n\<in>nat |] ==> intrel `` {<m,n>} : int"
  17.144 +by (simp add: int_def)
  17.145 +
  17.146 +declare equiv_intrel [THEN eq_equiv_class_iff, simp]
  17.147 +declare conj_cong [cong]
  17.148 +
  17.149 +lemmas eq_intrelD = eq_equiv_class [OF _ equiv_intrel]
  17.150 +
  17.151 +(** int_of: the injection from nat to int **)
  17.152 +
  17.153 +lemma int_of_type [simp,TC]: "$#m : int"
  17.154 +by (simp add: int_def quotient_def int_of_def, auto)
  17.155 +
  17.156 +lemma int_of_eq [iff]: "($# m = $# n) <-> natify(m)=natify(n)"
  17.157 +by (simp add: int_of_def)
  17.158 +
  17.159 +lemma int_of_inject: "[| $#m = $#n;  m\<in>nat;  n\<in>nat |] ==> m=n"
  17.160 +by (drule int_of_eq [THEN iffD1], auto)
  17.161 +
  17.162 +
  17.163 +(** intify: coercion from anything to int **)
  17.164 +
  17.165 +lemma intify_in_int [iff,TC]: "intify(x) : int"
  17.166 +by (simp add: intify_def)
  17.167 +
  17.168 +lemma intify_ident [simp]: "n : int ==> intify(n) = n"
  17.169 +by (simp add: intify_def)
  17.170 +
  17.171 +
  17.172 +subsection{*Collapsing rules: to remove @{term intify}
  17.173 +            from arithmetic expressions*}
  17.174 +
  17.175 +lemma intify_idem [simp]: "intify(intify(x)) = intify(x)"
  17.176 +by simp
  17.177 +
  17.178 +lemma int_of_natify [simp]: "$# (natify(m)) = $# m"
  17.179 +by (simp add: int_of_def)
  17.180 +
  17.181 +lemma zminus_intify [simp]: "$- (intify(m)) = $- m"
  17.182 +by (simp add: zminus_def)
  17.183 +
  17.184 +(** Addition **)
  17.185 +
  17.186 +lemma zadd_intify1 [simp]: "intify(x) $+ y = x $+ y"
  17.187 +by (simp add: zadd_def)
  17.188 +
  17.189 +lemma zadd_intify2 [simp]: "x $+ intify(y) = x $+ y"
  17.190 +by (simp add: zadd_def)
  17.191 +
  17.192 +(** Subtraction **)
  17.193 +
  17.194 +lemma zdiff_intify1 [simp]:"intify(x) $- y = x $- y"
  17.195 +by (simp add: zdiff_def)
  17.196 +
  17.197 +lemma zdiff_intify2 [simp]:"x $- intify(y) = x $- y"
  17.198 +by (simp add: zdiff_def)
  17.199 +
  17.200 +(** Multiplication **)
  17.201 +
  17.202 +lemma zmult_intify1 [simp]:"intify(x) $* y = x $* y"
  17.203 +by (simp add: zmult_def)
  17.204 +
  17.205 +lemma zmult_intify2 [simp]:"x $* intify(y) = x $* y"
  17.206 +by (simp add: zmult_def)
  17.207 +
  17.208 +(** Orderings **)
  17.209 +
  17.210 +lemma zless_intify1 [simp]:"intify(x) $< y <-> x $< y"
  17.211 +by (simp add: zless_def)
  17.212 +
  17.213 +lemma zless_intify2 [simp]:"x $< intify(y) <-> x $< y"
  17.214 +by (simp add: zless_def)
  17.215 +
  17.216 +lemma zle_intify1 [simp]:"intify(x) $<= y <-> x $<= y"
  17.217 +by (simp add: zle_def)
  17.218 +
  17.219 +lemma zle_intify2 [simp]:"x $<= intify(y) <-> x $<= y"
  17.220 +by (simp add: zle_def)
  17.221 +
  17.222 +
  17.223 +subsection{*@{term zminus}: unary negation on @{term int}*}
  17.224 +
  17.225 +lemma zminus_congruent: "(%<x,y>. intrel``{<y,x>}) respects intrel"
  17.226 +by (auto simp add: congruent_def add_ac)
  17.227 +
  17.228 +lemma raw_zminus_type: "z : int ==> raw_zminus(z) : int"
  17.229 +apply (simp add: int_def raw_zminus_def)
  17.230 +apply (typecheck add: UN_equiv_class_type [OF equiv_intrel zminus_congruent])
  17.231 +done
  17.232 +
  17.233 +lemma zminus_type [TC,iff]: "$-z : int"
  17.234 +by (simp add: zminus_def raw_zminus_type)
  17.235 +
  17.236 +lemma raw_zminus_inject: 
  17.237 +     "[| raw_zminus(z) = raw_zminus(w);  z: int;  w: int |] ==> z=w"
  17.238 +apply (simp add: int_def raw_zminus_def)
  17.239 +apply (erule UN_equiv_class_inject [OF equiv_intrel zminus_congruent], safe)
  17.240 +apply (auto dest: eq_intrelD simp add: add_ac)
  17.241 +done
  17.242 +
  17.243 +lemma zminus_inject_intify [dest!]: "$-z = $-w ==> intify(z) = intify(w)"
  17.244 +apply (simp add: zminus_def)
  17.245 +apply (blast dest!: raw_zminus_inject)
  17.246 +done
  17.247 +
  17.248 +lemma zminus_inject: "[| $-z = $-w;  z: int;  w: int |] ==> z=w"
  17.249 +by auto
  17.250 +
  17.251 +lemma raw_zminus: 
  17.252 +    "[| x\<in>nat;  y\<in>nat |] ==> raw_zminus(intrel``{<x,y>}) = intrel `` {<y,x>}"
  17.253 +apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent])
  17.254 +done
  17.255 +
  17.256 +lemma zminus: 
  17.257 +    "[| x\<in>nat;  y\<in>nat |]  
  17.258 +     ==> $- (intrel``{<x,y>}) = intrel `` {<y,x>}"
  17.259 +by (simp add: zminus_def raw_zminus image_intrel_int)
  17.260 +
  17.261 +lemma raw_zminus_zminus: "z : int ==> raw_zminus (raw_zminus(z)) = z"
  17.262 +by (auto simp add: int_def raw_zminus)
  17.263 +
  17.264 +lemma zminus_zminus_intify [simp]: "$- ($- z) = intify(z)"
  17.265 +by (simp add: zminus_def raw_zminus_type raw_zminus_zminus)
  17.266 +
  17.267 +lemma zminus_int0 [simp]: "$- ($#0) = $#0"
  17.268 +by (simp add: int_of_def zminus)
  17.269 +
  17.270 +lemma zminus_zminus: "z : int ==> $- ($- z) = z"
  17.271 +by simp
  17.272 +
  17.273 +
  17.274 +subsection{*@{term znegative}: the test for negative integers*}
  17.275 +
  17.276 +lemma znegative: "[| x\<in>nat; y\<in>nat |] ==> znegative(intrel``{<x,y>}) <-> x<y"
  17.277 +apply (cases "x<y") 
  17.278 +apply (auto simp add: znegative_def not_lt_iff_le)
  17.279 +apply (subgoal_tac "y #+ x2 < x #+ y2", force) 
  17.280 +apply (rule add_le_lt_mono, auto) 
  17.281 +done
  17.282 +
  17.283 +(*No natural number is negative!*)
  17.284 +lemma not_znegative_int_of [iff]: "~ znegative($# n)"
  17.285 +by (simp add: znegative int_of_def) 
  17.286 +
  17.287 +lemma znegative_zminus_int_of [simp]: "znegative($- $# succ(n))"
  17.288 +by (simp add: znegative int_of_def zminus natify_succ)
  17.289 +
  17.290 +lemma not_znegative_imp_zero: "~ znegative($- $# n) ==> natify(n)=0"
  17.291 +by (simp add: znegative int_of_def zminus Ord_0_lt_iff [THEN iff_sym])
  17.292 +
  17.293 +
  17.294 +subsection{*@{term nat_of}: Coercion of an Integer to a Natural Number*}
  17.295 +
  17.296 +lemma nat_of_intify [simp]: "nat_of(intify(z)) = nat_of(z)"
  17.297 +by (simp add: nat_of_def)
  17.298 +
  17.299 +lemma nat_of_congruent: "(\<lambda>x. (\<lambda>\<langle>x,y\<rangle>. x #- y)(x)) respects intrel"
  17.300 +by (auto simp add: congruent_def split add: nat_diff_split)
  17.301 +
  17.302 +lemma raw_nat_of: 
  17.303 +    "[| x\<in>nat;  y\<in>nat |] ==> raw_nat_of(intrel``{<x,y>}) = x#-y"
  17.304 +by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent])
  17.305 +
  17.306 +lemma raw_nat_of_int_of: "raw_nat_of($# n) = natify(n)"
  17.307 +by (simp add: int_of_def raw_nat_of)
  17.308 +
  17.309 +lemma nat_of_int_of [simp]: "nat_of($# n) = natify(n)"
  17.310 +by (simp add: raw_nat_of_int_of nat_of_def)
  17.311 +
  17.312 +lemma raw_nat_of_type: "raw_nat_of(z) \<in> nat"
  17.313 +by (simp add: raw_nat_of_def)
  17.314 +
  17.315 +lemma nat_of_type [iff,TC]: "nat_of(z) \<in> nat"
  17.316 +by (simp add: nat_of_def raw_nat_of_type)
  17.317 +
  17.318 +subsection{*zmagnitude: magnitide of an integer, as a natural number*}
  17.319 +
  17.320 +lemma zmagnitude_int_of [simp]: "zmagnitude($# n) = natify(n)"
  17.321 +by (auto simp add: zmagnitude_def int_of_eq)
  17.322 +
  17.323 +lemma natify_int_of_eq: "natify(x)=n ==> $#x = $# n"
  17.324 +apply (drule sym)
  17.325 +apply (simp (no_asm_simp) add: int_of_eq)
  17.326 +done
  17.327 +
  17.328 +lemma zmagnitude_zminus_int_of [simp]: "zmagnitude($- $# n) = natify(n)"
  17.329 +apply (simp add: zmagnitude_def)
  17.330 +apply (rule the_equality)
  17.331 +apply (auto dest!: not_znegative_imp_zero natify_int_of_eq
  17.332 +            iff del: int_of_eq, auto)
  17.333 +done
  17.334 +
  17.335 +lemma zmagnitude_type [iff,TC]: "zmagnitude(z)\<in>nat"
  17.336 +apply (simp add: zmagnitude_def)
  17.337 +apply (rule theI2, auto)
  17.338 +done
  17.339 +
  17.340 +lemma not_zneg_int_of: 
  17.341 +     "[| z: int; ~ znegative(z) |] ==> \<exists>n\<in>nat. z = $# n"
  17.342 +apply (auto simp add: int_def znegative int_of_def not_lt_iff_le)
  17.343 +apply (rename_tac x y) 
  17.344 +apply (rule_tac x="x#-y" in bexI) 
  17.345 +apply (auto simp add: add_diff_inverse2) 
  17.346 +done
  17.347 +
  17.348 +lemma not_zneg_mag [simp]:
  17.349 +     "[| z: int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z"
  17.350 +by (drule not_zneg_int_of, auto)
  17.351 +
  17.352 +lemma zneg_int_of: 
  17.353 +     "[| znegative(z); z: int |] ==> \<exists>n\<in>nat. z = $- ($# succ(n))"
  17.354 +by (auto simp add: int_def znegative zminus int_of_def dest!: less_imp_succ_add)
  17.355 +
  17.356 +lemma zneg_mag [simp]:
  17.357 +     "[| znegative(z); z: int |] ==> $# (zmagnitude(z)) = $- z"
  17.358 +by (drule zneg_int_of, auto)
  17.359 +
  17.360 +lemma int_cases: "z : int ==> \<exists>n\<in>nat. z = $# n | z = $- ($# succ(n))"
  17.361 +apply (case_tac "znegative (z) ")
  17.362 +prefer 2 apply (blast dest: not_zneg_mag sym)
  17.363 +apply (blast dest: zneg_int_of)
  17.364 +done
  17.365 +
  17.366 +lemma not_zneg_raw_nat_of:
  17.367 +     "[| ~ znegative(z); z: int |] ==> $# (raw_nat_of(z)) = z"
  17.368 +apply (drule not_zneg_int_of)
  17.369 +apply (auto simp add: raw_nat_of_type raw_nat_of_int_of)
  17.370 +done
  17.371 +
  17.372 +lemma not_zneg_nat_of_intify:
  17.373 +     "~ znegative(intify(z)) ==> $# (nat_of(z)) = intify(z)"
  17.374 +by (simp (no_asm_simp) add: nat_of_def not_zneg_raw_nat_of)
  17.375 +
  17.376 +lemma not_zneg_nat_of: "[| ~ znegative(z); z: int |] ==> $# (nat_of(z)) = z"
  17.377 +apply (simp (no_asm_simp) add: not_zneg_nat_of_intify)
  17.378 +done
  17.379 +
  17.380 +lemma zneg_nat_of [simp]: "znegative(intify(z)) ==> nat_of(z) = 0"
  17.381 +apply (subgoal_tac "intify(z) \<in> int")
  17.382 +apply (simp add: int_def) 
  17.383 +apply (auto simp add: znegative nat_of_def raw_nat_of 
  17.384 +            split add: nat_diff_split) 
  17.385 +done
  17.386 +
  17.387 +
  17.388 +subsection{*@{term zadd}: addition on int*}
  17.389 +
  17.390 +text{*Congruence Property for Addition*}
  17.391 +lemma zadd_congruent2: 
  17.392 +    "(%z1 z2. let <x1,y1>=z1; <x2,y2>=z2                  
  17.393 +                            in intrel``{<x1#+x2, y1#+y2>})
  17.394 +     respects2 intrel"
  17.395 +apply (simp add: congruent2_def)
  17.396 +(*Proof via congruent2_commuteI seems longer*)
  17.397 +apply safe
  17.398 +apply (simp (no_asm_simp) add: add_assoc Let_def)
  17.399 +(*The rest should be trivial, but rearranging terms is hard
  17.400 +  add_ac does not help rewriting with the assumptions.*)
  17.401 +apply (rule_tac m1 = x1a in add_left_commute [THEN ssubst])
  17.402 +apply (rule_tac m1 = x2a in add_left_commute [THEN ssubst])
  17.403 +apply (simp (no_asm_simp) add: add_assoc [symmetric])
  17.404 +done
  17.405 +
  17.406 +lemma raw_zadd_type: "[| z: int;  w: int |] ==> raw_zadd(z,w) : int"
  17.407 +apply (simp add: int_def raw_zadd_def)
  17.408 +apply (rule UN_equiv_class_type2 [OF equiv_intrel zadd_congruent2], assumption+)
  17.409 +apply (simp add: Let_def)
  17.410 +done
  17.411 +
  17.412 +lemma zadd_type [iff,TC]: "z $+ w : int"
  17.413 +by (simp add: zadd_def raw_zadd_type)
  17.414 +
  17.415 +lemma raw_zadd: 
  17.416 +  "[| x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat |]               
  17.417 +   ==> raw_zadd (intrel``{<x1,y1>}, intrel``{<x2,y2>}) =   
  17.418 +       intrel `` {<x1#+x2, y1#+y2>}"
  17.419 +apply (simp add: raw_zadd_def 
  17.420 +             UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2])
  17.421 +apply (simp add: Let_def)
  17.422 +done
  17.423 +
  17.424 +lemma zadd: 
  17.425 +  "[| x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat |]          
  17.426 +   ==> (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) =   
  17.427 +       intrel `` {<x1#+x2, y1#+y2>}"
  17.428 +by (simp add: zadd_def raw_zadd image_intrel_int)
  17.429 +
  17.430 +lemma raw_zadd_int0: "z : int ==> raw_zadd ($#0,z) = z"
  17.431 +by (auto simp add: int_def int_of_def raw_zadd)
  17.432 +
  17.433 +lemma zadd_int0_intify [simp]: "$#0 $+ z = intify(z)"
  17.434 +by (simp add: zadd_def raw_zadd_int0)
  17.435 +
  17.436 +lemma zadd_int0: "z: int ==> $#0 $+ z = z"
  17.437 +by simp
  17.438 +
  17.439 +lemma raw_zminus_zadd_distrib: 
  17.440 +     "[| z: int;  w: int |] ==> $- raw_zadd(z,w) = raw_zadd($- z, $- w)"
  17.441 +by (auto simp add: zminus raw_zadd int_def)
  17.442 +
  17.443 +lemma zminus_zadd_distrib [simp]: "$- (z $+ w) = $- z $+ $- w"
  17.444 +by (simp add: zadd_def raw_zminus_zadd_distrib)
  17.445 +
  17.446 +lemma raw_zadd_commute:
  17.447 +     "[| z: int;  w: int |] ==> raw_zadd(z,w) = raw_zadd(w,z)"
  17.448 +by (auto simp add: raw_zadd add_ac int_def)
  17.449 +
  17.450 +lemma zadd_commute: "z $+ w = w $+ z"
  17.451 +by (simp add: zadd_def raw_zadd_commute)
  17.452 +
  17.453 +lemma raw_zadd_assoc: 
  17.454 +    "[| z1: int;  z2: int;  z3: int |]    
  17.455 +     ==> raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))"
  17.456 +by (auto simp add: int_def raw_zadd add_assoc)
  17.457 +
  17.458 +lemma zadd_assoc: "(z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)"
  17.459 +by (simp add: zadd_def raw_zadd_type raw_zadd_assoc)
  17.460 +
  17.461 +(*For AC rewriting*)
  17.462 +lemma zadd_left_commute: "z1$+(z2$+z3) = z2$+(z1$+z3)"
  17.463 +apply (simp add: zadd_assoc [symmetric])
  17.464 +apply (simp add: zadd_commute)
  17.465 +done
  17.466 +
  17.467 +(*Integer addition is an AC operator*)
  17.468 +lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
  17.469 +
  17.470 +lemma int_of_add: "$# (m #+ n) = ($#m) $+ ($#n)"
  17.471 +by (simp add: int_of_def zadd)
  17.472 +
  17.473 +lemma int_succ_int_1: "$# succ(m) = $# 1 $+ ($# m)"
  17.474 +by (simp add: int_of_add [symmetric] natify_succ)
  17.475 +
  17.476 +lemma int_of_diff: 
  17.477 +     "[| m\<in>nat;  n le m |] ==> $# (m #- n) = ($#m) $- ($#n)"
  17.478 +apply (simp add: int_of_def zdiff_def)
  17.479 +apply (frule lt_nat_in_nat)
  17.480 +apply (simp_all add: zadd zminus add_diff_inverse2)
  17.481 +done
  17.482 +
  17.483 +lemma raw_zadd_zminus_inverse: "z : int ==> raw_zadd (z, $- z) = $#0"
  17.484 +by (auto simp add: int_def int_of_def zminus raw_zadd add_commute)
  17.485 +
  17.486 +lemma zadd_zminus_inverse [simp]: "z $+ ($- z) = $#0"
  17.487 +apply (simp add: zadd_def)
  17.488 +apply (subst zminus_intify [symmetric])
  17.489 +apply (rule intify_in_int [THEN raw_zadd_zminus_inverse])
  17.490 +done
  17.491 +
  17.492 +lemma zadd_zminus_inverse2 [simp]: "($- z) $+ z = $#0"
  17.493 +by (simp add: zadd_commute zadd_zminus_inverse)
  17.494 +
  17.495 +lemma zadd_int0_right_intify [simp]: "z $+ $#0 = intify(z)"
  17.496 +by (rule trans [OF zadd_commute zadd_int0_intify])
  17.497 +
  17.498 +lemma zadd_int0_right: "z:int ==> z $+ $#0 = z"
  17.499 +by simp
  17.500 +
  17.501 +
  17.502 +subsection{*@{term zmult}: Integer Multiplication*}
  17.503 +
  17.504 +text{*Congruence property for multiplication*}
  17.505 +lemma zmult_congruent2:
  17.506 +    "(%p1 p2. split(%x1 y1. split(%x2 y2.      
  17.507 +                    intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1))
  17.508 +     respects2 intrel"
  17.509 +apply (rule equiv_intrel [THEN congruent2_commuteI], auto)
  17.510 +(*Proof that zmult is congruent in one argument*)
  17.511 +apply (rename_tac x y)
  17.512 +apply (frule_tac t = "%u. x#*u" in sym [THEN subst_context])
  17.513 +apply (drule_tac t = "%u. y#*u" in subst_context)
  17.514 +apply (erule add_left_cancel)+
  17.515 +apply (simp_all add: add_mult_distrib_left)
  17.516 +done
  17.517 +
  17.518 +
  17.519 +lemma raw_zmult_type: "[| z: int;  w: int |] ==> raw_zmult(z,w) : int"
  17.520 +apply (simp add: int_def raw_zmult_def)
  17.521 +apply (rule UN_equiv_class_type2 [OF equiv_intrel zmult_congruent2], assumption+)
  17.522 +apply (simp add: Let_def)
  17.523 +done
  17.524 +
  17.525 +lemma zmult_type [iff,TC]: "z $* w : int"
  17.526 +by (simp add: zmult_def raw_zmult_type)
  17.527 +
  17.528 +lemma raw_zmult: 
  17.529 +     "[| x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat |]     
  17.530 +      ==> raw_zmult(intrel``{<x1,y1>}, intrel``{<x2,y2>}) =      
  17.531 +          intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}"
  17.532 +by (simp add: raw_zmult_def 
  17.533 +           UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2])
  17.534 +
  17.535 +lemma zmult: 
  17.536 +     "[| x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat |]     
  17.537 +      ==> (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) =      
  17.538 +          intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}"
  17.539 +by (simp add: zmult_def raw_zmult image_intrel_int)
  17.540 +
  17.541 +lemma raw_zmult_int0: "z : int ==> raw_zmult ($#0,z) = $#0"
  17.542 +by (auto simp add: int_def int_of_def raw_zmult)
  17.543 +
  17.544 +lemma zmult_int0 [simp]: "$#0 $* z = $#0"
  17.545 +by (simp add: zmult_def raw_zmult_int0)
  17.546 +
  17.547 +lemma raw_zmult_int1: "z : int ==> raw_zmult ($#1,z) = z"
  17.548 +by (auto simp add: int_def int_of_def raw_zmult)
  17.549 +
  17.550 +lemma zmult_int1_intify [simp]: "$#1 $* z = intify(z)"
  17.551 +by (simp add: zmult_def raw_zmult_int1)
  17.552 +
  17.553 +lemma zmult_int1: "z : int ==> $#1 $* z = z"
  17.554 +by simp
  17.555 +
  17.556 +lemma raw_zmult_commute:
  17.557 +     "[| z: int;  w: int |] ==> raw_zmult(z,w) = raw_zmult(w,z)"
  17.558 +by (auto simp add: int_def raw_zmult add_ac mult_ac)
  17.559 +
  17.560 +lemma zmult_commute: "z $* w = w $* z"
  17.561 +by (simp add: zmult_def raw_zmult_commute)
  17.562 +
  17.563 +lemma raw_zmult_zminus: 
  17.564 +     "[| z: int;  w: int |] ==> raw_zmult($- z, w) = $- raw_zmult(z, w)"
  17.565 +by (auto simp add: int_def zminus raw_zmult add_ac)
  17.566 +
  17.567 +lemma zmult_zminus [simp]: "($- z) $* w = $- (z $* w)"
  17.568 +apply (simp add: zmult_def raw_zmult_zminus)
  17.569 +apply (subst zminus_intify [symmetric], rule raw_zmult_zminus, auto)
  17.570 +done
  17.571 +
  17.572 +lemma zmult_zminus_right [simp]: "w $* ($- z) = $- (w $* z)"
  17.573 +by (simp add: zmult_commute [of w])
  17.574 +
  17.575 +lemma raw_zmult_assoc: 
  17.576 +    "[| z1: int;  z2: int;  z3: int |]    
  17.577 +     ==> raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))"
  17.578 +by (auto simp add: int_def raw_zmult add_mult_distrib_left add_ac mult_ac)
  17.579 +
  17.580 +lemma zmult_assoc: "(z1 $* z2) $* z3 = z1 $* (z2 $* z3)"
  17.581 +by (simp add: zmult_def raw_zmult_type raw_zmult_assoc)
  17.582 +
  17.583 +(*For AC rewriting*)
  17.584 +lemma zmult_left_commute: "z1$*(z2$*z3) = z2$*(z1$*z3)"
  17.585 +apply (simp add: zmult_assoc [symmetric])
  17.586 +apply (simp add: zmult_commute)
  17.587 +done
  17.588 +
  17.589 +(*Integer multiplication is an AC operator*)
  17.590 +lemmas zmult_ac = zmult_assoc zmult_commute zmult_left_commute
  17.591 +
  17.592 +lemma raw_zadd_zmult_distrib: 
  17.593 +    "[| z1: int;  z2: int;  w: int |]   
  17.594 +     ==> raw_zmult(raw_zadd(z1,z2), w) =  
  17.595 +         raw_zadd (raw_zmult(z1,w), raw_zmult(z2,w))"
  17.596 +by (auto simp add: int_def raw_zadd raw_zmult add_mult_distrib_left add_ac mult_ac)
  17.597 +
  17.598 +lemma zadd_zmult_distrib: "(z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)"
  17.599 +by (simp add: zmult_def zadd_def raw_zadd_type raw_zmult_type 
  17.600 +              raw_zadd_zmult_distrib)
  17.601 +
  17.602 +lemma zadd_zmult_distrib2: "w $* (z1 $+ z2) = (w $* z1) $+ (w $* z2)"
  17.603 +by (simp add: zmult_commute [of w] zadd_zmult_distrib)
  17.604 +
  17.605 +lemmas int_typechecks = 
  17.606 +  int_of_type zminus_type zmagnitude_type zadd_type zmult_type
  17.607 +
  17.608 +
  17.609 +(*** Subtraction laws ***)
  17.610 +
  17.611 +lemma zdiff_type [iff,TC]: "z $- w : int"
  17.612 +by (simp add: zdiff_def)
  17.613 +
  17.614 +lemma zminus_zdiff_eq [simp]: "$- (z $- y) = y $- z"
  17.615 +by (simp add: zdiff_def zadd_commute)
  17.616 +
  17.617 +lemma zdiff_zmult_distrib: "(z1 $- z2) $* w = (z1 $* w) $- (z2 $* w)"
  17.618 +apply (simp add: zdiff_def)
  17.619 +apply (subst zadd_zmult_distrib)
  17.620 +apply (simp add: zmult_zminus)
  17.621 +done
  17.622 +
  17.623 +lemma zdiff_zmult_distrib2: "w $* (z1 $- z2) = (w $* z1) $- (w $* z2)"
  17.624 +by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
  17.625 +
  17.626 +lemma zadd_zdiff_eq: "x $+ (y $- z) = (x $+ y) $- z"
  17.627 +by (simp add: zdiff_def zadd_ac)
  17.628 +
  17.629 +lemma zdiff_zadd_eq: "(x $- y) $+ z = (x $+ z) $- y"
  17.630 +by (simp add: zdiff_def zadd_ac)
  17.631 +
  17.632 +
  17.633 +subsection{*The "Less Than" Relation*}
  17.634 +
  17.635 +(*"Less than" is a linear ordering*)
  17.636 +lemma zless_linear_lemma: 
  17.637 +     "[| z: int; w: int |] ==> z$<w | z=w | w$<z"
  17.638 +apply (simp add: int_def zless_def znegative_def zdiff_def, auto)
  17.639 +apply (simp add: zadd zminus image_iff Bex_def)
  17.640 +apply (rule_tac i = "xb#+ya" and j = "xc #+ y" in Ord_linear_lt)
  17.641 +apply (force dest!: spec simp add: add_ac)+
  17.642 +done
  17.643 +
  17.644 +lemma zless_linear: "z$<w | intify(z)=intify(w) | w$<z"
  17.645 +apply (cut_tac z = " intify (z) " and w = " intify (w) " in zless_linear_lemma)
  17.646 +apply auto
  17.647 +done
  17.648 +
  17.649 +lemma zless_not_refl [iff]: "~ (z$<z)"
  17.650 +by (auto simp add: zless_def znegative_def int_of_def zdiff_def)
  17.651 +
  17.652 +lemma neq_iff_zless: "[| x: int; y: int |] ==> (x ~= y) <-> (x $< y | y $< x)"
  17.653 +by (cut_tac z = x and w = y in zless_linear, auto)
  17.654 +
  17.655 +lemma zless_imp_intify_neq: "w $< z ==> intify(w) ~= intify(z)"
  17.656 +apply auto
  17.657 +apply (subgoal_tac "~ (intify (w) $< intify (z))")
  17.658 +apply (erule_tac [2] ssubst)
  17.659 +apply (simp (no_asm_use))
  17.660 +apply auto
  17.661 +done
  17.662 +
  17.663 +(*This lemma allows direct proofs of other <-properties*)
  17.664 +lemma zless_imp_succ_zadd_lemma: 
  17.665 +    "[| w $< z; w: int; z: int |] ==> (\<exists>n\<in>nat. z = w $+ $#(succ(n)))"
  17.666 +apply (simp add: zless_def znegative_def zdiff_def int_def)
  17.667 +apply (auto dest!: less_imp_succ_add simp add: zadd zminus int_of_def)
  17.668 +apply (rule_tac x = k in bexI)
  17.669 +apply (erule add_left_cancel, auto)
  17.670 +done
  17.671 +
  17.672 +lemma zless_imp_succ_zadd:
  17.673 +     "w $< z ==> (\<exists>n\<in>nat. w $+ $#(succ(n)) = intify(z))"
  17.674 +apply (subgoal_tac "intify (w) $< intify (z) ")
  17.675 +apply (drule_tac w = "intify (w) " in zless_imp_succ_zadd_lemma)
  17.676 +apply auto
  17.677 +done
  17.678 +
  17.679 +lemma zless_succ_zadd_lemma: 
  17.680 +    "w : int ==> w $< w $+ $# succ(n)"
  17.681 +apply (simp add: zless_def znegative_def zdiff_def int_def)
  17.682 +apply (auto simp add: zadd zminus int_of_def image_iff)
  17.683 +apply (rule_tac x = 0 in exI, auto)
  17.684 +done
  17.685 +
  17.686 +lemma zless_succ_zadd: "w $< w $+ $# succ(n)"
  17.687 +by (cut_tac intify_in_int [THEN zless_succ_zadd_lemma], auto)
  17.688 +
  17.689 +lemma zless_iff_succ_zadd:
  17.690 +     "w $< z <-> (\<exists>n\<in>nat. w $+ $#(succ(n)) = intify(z))"
  17.691 +apply (rule iffI)
  17.692 +apply (erule zless_imp_succ_zadd, auto)
  17.693 +apply (rename_tac "n")
  17.694 +apply (cut_tac w = w and n = n in zless_succ_zadd, auto)
  17.695 +done
  17.696 +
  17.697 +lemma zless_int_of [simp]: "[| m\<in>nat; n\<in>nat |] ==> ($#m $< $#n) <-> (m<n)"
  17.698 +apply (simp add: less_iff_succ_add zless_iff_succ_zadd int_of_add [symmetric])
  17.699 +apply (blast intro: sym)
  17.700 +done
  17.701 +
  17.702 +lemma zless_trans_lemma: 
  17.703 +    "[| x $< y; y $< z; x: int; y : int; z: int |] ==> x $< z"
  17.704 +apply (simp add: zless_def znegative_def zdiff_def int_def)
  17.705 +apply (auto simp add: zadd zminus image_iff)
  17.706 +apply (rename_tac x1 x2 y1 y2)
  17.707 +apply (rule_tac x = "x1#+x2" in exI)
  17.708 +apply (rule_tac x = "y1#+y2" in exI)
  17.709 +apply (auto simp add: add_lt_mono)
  17.710 +apply (rule sym)
  17.711 +apply (erule add_left_cancel)+
  17.712 +apply auto
  17.713 +done
  17.714 +
  17.715 +lemma zless_trans: "[| x $< y; y $< z |] ==> x $< z"
  17.716 +apply (subgoal_tac "intify (x) $< intify (z) ")
  17.717 +apply (rule_tac [2] y = "intify (y) " in zless_trans_lemma)
  17.718 +apply auto
  17.719 +done
  17.720 +
  17.721 +lemma zless_not_sym: "z $< w ==> ~ (w $< z)"
  17.722 +by (blast dest: zless_trans)
  17.723 +
  17.724 +(* [| z $< w; ~ P ==> w $< z |] ==> P *)
  17.725 +lemmas zless_asym = zless_not_sym [THEN swap, standard]
  17.726 +
  17.727 +lemma zless_imp_zle: "z $< w ==> z $<= w"
  17.728 +by (simp add: zle_def)
  17.729 +
  17.730 +lemma zle_linear: "z $<= w | w $<= z"
  17.731 +apply (simp add: zle_def)
  17.732 +apply (cut_tac zless_linear, blast)
  17.733 +done
  17.734 +
  17.735 +
  17.736 +subsection{*Less Than or Equals*}
  17.737 +
  17.738 +lemma zle_refl: "z $<= z"
  17.739 +by (simp add: zle_def)
  17.740 +
  17.741 +lemma zle_eq_refl: "x=y ==> x $<= y"
  17.742 +by (simp add: zle_refl)
  17.743 +
  17.744 +lemma zle_anti_sym_intify: "[| x $<= y; y $<= x |] ==> intify(x) = intify(y)"
  17.745 +apply (simp add: zle_def, auto)
  17.746 +apply (blast dest: zless_trans)
  17.747 +done
  17.748 +
  17.749 +lemma zle_anti_sym: "[| x $<= y; y $<= x; x: int; y: int |] ==> x=y"
  17.750 +by (drule zle_anti_sym_intify, auto)
  17.751 +
  17.752 +lemma zle_trans_lemma:
  17.753 +     "[| x: int; y: int; z: int; x $<= y; y $<= z |] ==> x $<= z"
  17.754 +apply (simp add: zle_def, auto)
  17.755 +apply (blast intro: zless_trans)
  17.756 +done
  17.757 +
  17.758 +lemma zle_trans: "[| x $<= y; y $<= z |] ==> x $<= z"
  17.759 +apply (subgoal_tac "intify (x) $<= intify (z) ")
  17.760 +apply (rule_tac [2] y = "intify (y) " in zle_trans_lemma)
  17.761 +apply auto
  17.762 +done
  17.763 +
  17.764 +lemma zle_zless_trans: "[| i $<= j; j $< k |] ==> i $< k"
  17.765 +apply (auto simp add: zle_def)
  17.766 +apply (blast intro: zless_trans)
  17.767 +apply (simp add: zless_def zdiff_def zadd_def)
  17.768 +done
  17.769 +
  17.770 +lemma zless_zle_trans: "[| i $< j; j $<= k |] ==> i $< k"
  17.771 +apply (auto simp add: zle_def)
  17.772 +apply (blast intro: zless_trans)
  17.773 +apply (simp add: zless_def zdiff_def zminus_def)
  17.774 +done
  17.775 +
  17.776 +lemma not_zless_iff_zle: "~ (z $< w) <-> (w $<= z)"
  17.777 +apply (cut_tac z = z and w = w in zless_linear)
  17.778 +apply (auto dest: zless_trans simp add: zle_def)
  17.779 +apply (auto dest!: zless_imp_intify_neq)
  17.780 +done
  17.781 +
  17.782 +lemma not_zle_iff_zless: "~ (z $<= w) <-> (w $< z)"
  17.783 +by (simp add: not_zless_iff_zle [THEN iff_sym])
  17.784 +
  17.785 +
  17.786 +subsection{*More subtraction laws (for @{text zcompare_rls})*}
  17.787 +
  17.788 +lemma zdiff_zdiff_eq: "(x $- y) $- z = x $- (y $+ z)"
  17.789 +by (simp add: zdiff_def zadd_ac)
  17.790 +
  17.791 +lemma zdiff_zdiff_eq2: "x $- (y $- z) = (x $+ z) $- y"
  17.792 +by (simp add: zdiff_def zadd_ac)
  17.793 +
  17.794 +lemma zdiff_zless_iff: "(x$-y $< z) <-> (x $< z $+ y)"
  17.795 +by (simp add: zless_def zdiff_def zadd_ac)
  17.796 +
  17.797 +lemma zless_zdiff_iff: "(x $< z$-y) <-> (x $+ y $< z)"
  17.798 +by (simp add: zless_def zdiff_def zadd_ac)
  17.799 +
  17.800 +lemma zdiff_eq_iff: "[| x: int; z: int |] ==> (x$-y = z) <-> (x = z $+ y)"
  17.801 +by (auto simp add: zdiff_def zadd_assoc)
  17.802 +
  17.803 +lemma eq_zdiff_iff: "[| x: int; z: int |] ==> (x = z$-y) <-> (x $+ y = z)"
  17.804 +by (auto simp add: zdiff_def zadd_assoc)
  17.805 +
  17.806 +lemma zdiff_zle_iff_lemma:
  17.807 +     "[| x: int; z: int |] ==> (x$-y $<= z) <-> (x $<= z $+ y)"
  17.808 +by (auto simp add: zle_def zdiff_eq_iff zdiff_zless_iff)
  17.809 +
  17.810 +lemma zdiff_zle_iff: "(x$-y $<= z) <-> (x $<= z $+ y)"
  17.811 +by (cut_tac zdiff_zle_iff_lemma [OF intify_in_int intify_in_int], simp)
  17.812 +
  17.813 +lemma zle_zdiff_iff_lemma:
  17.814 +     "[| x: int; z: int |] ==>(x $<= z$-y) <-> (x $+ y $<= z)"
  17.815 +apply (auto simp add: zle_def zdiff_eq_iff zless_zdiff_iff)
  17.816 +apply (auto simp add: zdiff_def zadd_assoc)
  17.817 +done
  17.818 +
  17.819 +lemma zle_zdiff_iff: "(x $<= z$-y) <-> (x $+ y $<= z)"
  17.820 +by (cut_tac zle_zdiff_iff_lemma [ OF intify_in_int intify_in_int], simp)
  17.821 +
  17.822 +text{*This list of rewrites simplifies (in)equalities by bringing subtractions
  17.823 +  to the top and then moving negative terms to the other side.  
  17.824 +  Use with @{text zadd_ac}*}
  17.825 +lemmas zcompare_rls =
  17.826 +     zdiff_def [symmetric]
  17.827 +     zadd_zdiff_eq zdiff_zadd_eq zdiff_zdiff_eq zdiff_zdiff_eq2 
  17.828 +     zdiff_zless_iff zless_zdiff_iff zdiff_zle_iff zle_zdiff_iff 
  17.829 +     zdiff_eq_iff eq_zdiff_iff
  17.830 +
  17.831 +
  17.832 +subsection{*Monotonicity and Cancellation Results for Instantiation
  17.833 +     of the CancelNumerals Simprocs*}
  17.834 +
  17.835 +lemma zadd_left_cancel:
  17.836 +     "[| w: int; w': int |] ==> (z $+ w' = z $+ w) <-> (w' = w)"
  17.837 +apply safe
  17.838 +apply (drule_tac t = "%x. x $+ ($-z) " in subst_context)
  17.839 +apply (simp add: zadd_ac)
  17.840 +done
  17.841 +
  17.842 +lemma zadd_left_cancel_intify [simp]:
  17.843 +     "(z $+ w' = z $+ w) <-> intify(w') = intify(w)"
  17.844 +apply (rule iff_trans)
  17.845 +apply (rule_tac [2] zadd_left_cancel, auto)
  17.846 +done
  17.847 +
  17.848 +lemma zadd_right_cancel:
  17.849 +     "[| w: int; w': int |] ==> (w' $+ z = w $+ z) <-> (w' = w)"
  17.850 +apply safe
  17.851 +apply (drule_tac t = "%x. x $+ ($-z) " in subst_context)
  17.852 +apply (simp add: zadd_ac)
  17.853 +done
  17.854 +
  17.855 +lemma zadd_right_cancel_intify [simp]:
  17.856 +     "(w' $+ z = w $+ z) <-> intify(w') = intify(w)"
  17.857 +apply (rule iff_trans)
  17.858 +apply (rule_tac [2] zadd_right_cancel, auto)
  17.859 +done
  17.860 +
  17.861 +lemma zadd_right_cancel_zless [simp]: "(w' $+ z $< w $+ z) <-> (w' $< w)"
  17.862 +by (simp add: zdiff_zless_iff [THEN iff_sym] zdiff_def zadd_assoc)
  17.863 +
  17.864 +lemma zadd_left_cancel_zless [simp]: "(z $+ w' $< z $+ w) <-> (w' $< w)"
  17.865 +by (simp add: zadd_commute [of z] zadd_right_cancel_zless)
  17.866 +
  17.867 +lemma zadd_right_cancel_zle [simp]: "(w' $+ z $<= w $+ z) <-> w' $<= w"
  17.868 +by (simp add: zle_def)
  17.869 +
  17.870 +lemma zadd_left_cancel_zle [simp]: "(z $+ w' $<= z $+ w) <->  w' $<= w"
  17.871 +by (simp add: zadd_commute [of z]  zadd_right_cancel_zle)
  17.872 +
  17.873 +
  17.874 +(*"v $<= w ==> v$+z $<= w$+z"*)
  17.875 +lemmas zadd_zless_mono1 = zadd_right_cancel_zless [THEN iffD2, standard]
  17.876 +
  17.877 +(*"v $<= w ==> z$+v $<= z$+w"*)
  17.878 +lemmas zadd_zless_mono2 = zadd_left_cancel_zless [THEN iffD2, standard]
  17.879 +
  17.880 +(*"v $<= w ==> v$+z $<= w$+z"*)
  17.881 +lemmas zadd_zle_mono1 = zadd_right_cancel_zle [THEN iffD2, standard]
  17.882 +
  17.883 +(*"v $<= w ==> z$+v $<= z$+w"*)
  17.884 +lemmas zadd_zle_mono2 = zadd_left_cancel_zle [THEN iffD2, standard]
  17.885 +
  17.886 +lemma zadd_zle_mono: "[| w' $<= w; z' $<= z |] ==> w' $+ z' $<= w $+ z"
  17.887 +by (erule zadd_zle_mono1 [THEN zle_trans], simp)
  17.888 +
  17.889 +lemma zadd_zless_mono: "[| w' $< w; z' $<= z |] ==> w' $+ z' $< w $+ z"
  17.890 +by (erule zadd_zless_mono1 [THEN zless_zle_trans], simp)
  17.891 +
  17.892 +
  17.893 +subsection{*Comparison laws*}
  17.894 +
  17.895 +lemma zminus_zless_zminus [simp]: "($- x $< $- y) <-> (y $< x)"
  17.896 +by (simp add: zless_def zdiff_def zadd_ac)
  17.897 +
  17.898 +lemma zminus_zle_zminus [simp]: "($- x $<= $- y) <-> (y $<= x)"
  17.899 +by (simp add: not_zless_iff_zle [THEN iff_sym])
  17.900 +
  17.901 +subsubsection{*More inequality lemmas*}
  17.902 +
  17.903 +lemma equation_zminus: "[| x: int;  y: int |] ==> (x = $- y) <-> (y = $- x)"
  17.904 +by auto
  17.905 +
  17.906 +lemma zminus_equation: "[| x: int;  y: int |] ==> ($- x = y) <-> ($- y = x)"
  17.907 +by auto
  17.908 +
  17.909 +lemma equation_zminus_intify: "(intify(x) = $- y) <-> (intify(y) = $- x)"
  17.910 +apply (cut_tac x = "intify (x) " and y = "intify (y) " in equation_zminus)
  17.911 +apply auto
  17.912 +done
  17.913 +
  17.914 +lemma zminus_equation_intify: "($- x = intify(y)) <-> ($- y = intify(x))"
  17.915 +apply (cut_tac x = "intify (x) " and y = "intify (y) " in zminus_equation)
  17.916 +apply auto
  17.917 +done
  17.918 +
  17.919 +
  17.920 +subsubsection{*The next several equations are permutative: watch out!*}
  17.921 +
  17.922 +lemma zless_zminus: "(x $< $- y) <-> (y $< $- x)"
  17.923 +by (simp add: zless_def zdiff_def zadd_ac)
  17.924 +
  17.925 +lemma zminus_zless: "($- x $< y) <-> ($- y $< x)"
  17.926 +by (simp add: zless_def zdiff_def zadd_ac)
  17.927 +
  17.928 +lemma zle_zminus: "(x $<= $- y) <-> (y $<= $- x)"
  17.929 +by (simp add: not_zless_iff_zle [THEN iff_sym] zminus_zless)
  17.930 +
  17.931 +lemma zminus_zle: "($- x $<= y) <-> ($- y $<= x)"
  17.932 +by (simp add: not_zless_iff_zle [THEN iff_sym] zless_zminus)
  17.933 +
  17.934 +end
    18.1 --- a/src/ZF/IsaMakefile	Mon Feb 11 15:19:17 2008 +0100
    18.2 +++ b/src/ZF/IsaMakefile	Mon Feb 11 15:40:21 2008 +0100
    18.3 @@ -30,9 +30,9 @@
    18.4  
    18.5  $(OUT)/ZF$(ML_SUFFIX): $(OUT)/FOL$(ML_SUFFIX) AC.thy Arith.thy		\
    18.6    ArithSimp.thy Bin.thy Bool.thy Cardinal.thy CardinalArith.thy		\
    18.7 -  Cardinal_AC.thy Datatype.thy Epsilon.thy EquivClass.thy Finite.thy	\
    18.8 -  Fixedpt.thy Inductive.thy InfDatatype.thy Int.thy IntArith.thy	\
    18.9 -  IntDiv.thy List.thy Main.thy Main_ZFC.thy Nat.thy OrdQuant.thy	\
   18.10 +  Cardinal_AC.thy Datatype_ZF.thy Epsilon.thy EquivClass.thy Finite.thy	\
   18.11 +  Fixedpt.thy Inductive_ZF.thy InfDatatype.thy Int_ZF.thy IntArith.thy	\
   18.12 +  IntDiv_ZF.thy List_ZF.thy Main.thy Main_ZF.thy Main_ZFC.thy Nat_ZF.thy OrdQuant.thy	\
   18.13    Order.thy OrderArith.thy OrderType.thy Ordinal.thy Perm.thy QPair.thy	\
   18.14    QUniv.thy ROOT.ML Sum.thy Tools/cartprod.ML Tools/datatype_package.ML	\
   18.15    Tools/ind_cases.ML Tools/induct_tacs.ML Tools/inductive_package.ML	\
    19.1 --- a/src/ZF/List.thy	Mon Feb 11 15:19:17 2008 +0100
    19.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    19.3 @@ -1,1251 +0,0 @@
    19.4 -(*  Title:      ZF/List
    19.5 -    ID:         $Id$
    19.6 -    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    19.7 -    Copyright   1994  University of Cambridge
    19.8 -
    19.9 -*)
   19.10 -
   19.11 -header{*Lists in Zermelo-Fraenkel Set Theory*}
   19.12 -
   19.13 -theory List imports Datatype ArithSimp begin
   19.14 -
   19.15 -consts
   19.16 -  list       :: "i=>i"
   19.17 -
   19.18 -datatype
   19.19 -  "list(A)" = Nil | Cons ("a:A", "l: list(A)")
   19.20 -
   19.21 -
   19.22 -syntax
   19.23 - "[]"        :: i                                       ("[]")
   19.24 - "@List"     :: "is => i"                                 ("[(_)]")
   19.25 -
   19.26 -translations
   19.27 -  "[x, xs]"     == "Cons(x, [xs])"
   19.28 -  "[x]"         == "Cons(x, [])"
   19.29 -  "[]"          == "Nil"
   19.30 -
   19.31 -
   19.32 -consts
   19.33 -  length :: "i=>i"
   19.34 -  hd     :: "i=>i"
   19.35 -  tl     :: "i=>i"
   19.36 -
   19.37 -primrec
   19.38 -  "length([]) = 0"
   19.39 -  "length(Cons(a,l)) = succ(length(l))"
   19.40 -
   19.41 -primrec
   19.42 -  "hd([]) = 0"
   19.43 -  "hd(Cons(a,l)) = a"
   19.44 -
   19.45 -primrec
   19.46 -  "tl([]) = []"
   19.47 -  "tl(Cons(a,l)) = l"
   19.48 -
   19.49 -
   19.50 -consts
   19.51 -  map         :: "[i=>i, i] => i"
   19.52 -  set_of_list :: "i=>i"
   19.53 -  app         :: "[i,i]=>i"                        (infixr "@" 60)
   19.54 -
   19.55 -(*map is a binding operator -- it applies to meta-level functions, not
   19.56 -object-level functions.  This simplifies the final form of term_rec_conv,
   19.57 -although complicating its derivation.*)
   19.58 -primrec
   19.59 -  "map(f,[]) = []"
   19.60 -  "map(f,Cons(a,l)) = Cons(f(a), map(f,l))"
   19.61 -
   19.62 -primrec
   19.63 -  "set_of_list([]) = 0"
   19.64 -  "set_of_list(Cons(a,l)) = cons(a, set_of_list(l))"
   19.65 -
   19.66 -primrec
   19.67 -  app_Nil:  "[] @ ys = ys"
   19.68 -  app_Cons: "(Cons(a,l)) @ ys = Cons(a, l @ ys)"
   19.69 -
   19.70 -
   19.71 -consts
   19.72 -  rev :: "i=>i"
   19.73 -  flat       :: "i=>i"
   19.74 -  list_add   :: "i=>i"
   19.75 -
   19.76 -primrec
   19.77 -  "rev([]) = []"
   19.78 -  "rev(Cons(a,l)) = rev(l) @ [a]"
   19.79 -
   19.80 -primrec
   19.81 -  "flat([]) = []"
   19.82 -  "flat(Cons(l,ls)) = l @ flat(ls)"
   19.83 -
   19.84 -primrec
   19.85 -  "list_add([]) = 0"
   19.86 -  "list_add(Cons(a,l)) = a #+ list_add(l)"
   19.87 -
   19.88 -consts
   19.89 -  drop       :: "[i,i]=>i"
   19.90 -
   19.91 -primrec
   19.92 -  drop_0:    "drop(0,l) = l"
   19.93 -  drop_succ: "drop(succ(i), l) = tl (drop(i,l))"
   19.94 -
   19.95 -
   19.96 -(*** Thanks to Sidi Ehmety for the following ***)
   19.97 -
   19.98 -definition
   19.99 -(* Function `take' returns the first n elements of a list *)
  19.100 -  take     :: "[i,i]=>i"  where
  19.101 -  "take(n, as) == list_rec(lam n:nat. [],
  19.102 -		%a l r. lam n:nat. nat_case([], %m. Cons(a, r`m), n), as)`n"
  19.103 -
  19.104 -definition
  19.105 -  nth :: "[i, i]=>i"  where
  19.106 -  --{*returns the (n+1)th element of a list, or 0 if the
  19.107 -   list is too short.*}
  19.108 -  "nth(n, as) == list_rec(lam n:nat. 0,
  19.109 - 		          %a l r. lam n:nat. nat_case(a, %m. r`m, n), as) ` n"
  19.110 -
  19.111 -definition
  19.112 -  list_update :: "[i, i, i]=>i"  where
  19.113 -  "list_update(xs, i, v) == list_rec(lam n:nat. Nil,
  19.114 -      %u us vs. lam n:nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i"
  19.115 -
  19.116 -consts
  19.117 -  filter :: "[i=>o, i] => i"
  19.118 -  upt :: "[i, i] =>i"
  19.119 -
  19.120 -primrec
  19.121 -  "filter(P, Nil) = Nil"
  19.122 -  "filter(P, Cons(x, xs)) =
  19.123 -     (if P(x) then Cons(x, filter(P, xs)) else filter(P, xs))"
  19.124 -
  19.125 -primrec
  19.126 -  "upt(i, 0) = Nil"
  19.127 -  "upt(i, succ(j)) = (if i le j then upt(i, j)@[j] else Nil)"
  19.128 -
  19.129 -definition
  19.130 -  min :: "[i,i] =>i"  where
  19.131 -    "min(x, y) == (if x le y then x else y)"
  19.132 -
  19.133 -definition
  19.134 -  max :: "[i, i] =>i"  where
  19.135 -    "max(x, y) == (if x le y then y else x)"
  19.136 -
  19.137 -(*** Aspects of the datatype definition ***)
  19.138 -
  19.139 -declare list.intros [simp,TC]
  19.140 -
  19.141 -(*An elimination rule, for type-checking*)
  19.142 -inductive_cases ConsE: "Cons(a,l) : list(A)"
  19.143 -
  19.144 -lemma Cons_type_iff [simp]: "Cons(a,l) \<in> list(A) <-> a \<in> A & l \<in> list(A)"
  19.145 -by (blast elim: ConsE) 
  19.146 -
  19.147 -(*Proving freeness results*)
  19.148 -lemma Cons_iff: "Cons(a,l)=Cons(a',l') <-> a=a' & l=l'"
  19.149 -by auto
  19.150 -
  19.151 -lemma Nil_Cons_iff: "~ Nil=Cons(a,l)"
  19.152 -by auto
  19.153 -
  19.154 -lemma list_unfold: "list(A) = {0} + (A * list(A))"
  19.155 -by (blast intro!: list.intros [unfolded list.con_defs]
  19.156 -          elim: list.cases [unfolded list.con_defs])
  19.157 -
  19.158 -
  19.159 -(**  Lemmas to justify using "list" in other recursive type definitions **)
  19.160 -
  19.161 -lemma list_mono: "A<=B ==> list(A) <= list(B)"
  19.162 -apply (unfold list.defs )
  19.163 -apply (rule lfp_mono)
  19.164 -apply (simp_all add: list.bnd_mono)
  19.165 -apply (assumption | rule univ_mono basic_monos)+
  19.166 -done
  19.167 -
  19.168 -(*There is a similar proof by list induction.*)
  19.169 -lemma list_univ: "list(univ(A)) <= univ(A)"
  19.170 -apply (unfold list.defs list.con_defs)
  19.171 -apply (rule lfp_lowerbound)
  19.172 -apply (rule_tac [2] A_subset_univ [THEN univ_mono])
  19.173 -apply (blast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ)
  19.174 -done
  19.175 -
  19.176 -(*These two theorems justify datatypes involving list(nat), list(A), ...*)
  19.177 -lemmas list_subset_univ = subset_trans [OF list_mono list_univ]
  19.178 -
  19.179 -lemma list_into_univ: "[| l: list(A);  A <= univ(B) |] ==> l: univ(B)"
  19.180 -by (blast intro: list_subset_univ [THEN subsetD])
  19.181 -
  19.182 -lemma list_case_type:
  19.183 -    "[| l: list(A);
  19.184 -        c: C(Nil);
  19.185 -        !!x y. [| x: A;  y: list(A) |] ==> h(x,y): C(Cons(x,y))
  19.186 -     |] ==> list_case(c,h,l) : C(l)"
  19.187 -by (erule list.induct, auto)
  19.188 -
  19.189 -lemma list_0_triv: "list(0) = {Nil}"
  19.190 -apply (rule equalityI, auto) 
  19.191 -apply (induct_tac x, auto) 
  19.192 -done
  19.193 -
  19.194 -
  19.195 -(*** List functions ***)
  19.196 -
  19.197 -lemma tl_type: "l: list(A) ==> tl(l) : list(A)"
  19.198 -apply (induct_tac "l")
  19.199 -apply (simp_all (no_asm_simp) add: list.intros)
  19.200 -done
  19.201 -
  19.202 -(** drop **)
  19.203 -
  19.204 -lemma drop_Nil [simp]: "i:nat ==> drop(i, Nil) = Nil"
  19.205 -apply (induct_tac "i")
  19.206 -apply (simp_all (no_asm_simp))
  19.207 -done
  19.208 -
  19.209 -lemma drop_succ_Cons [simp]: "i:nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)"
  19.210 -apply (rule sym)
  19.211 -apply (induct_tac "i")
  19.212 -apply (simp (no_asm))
  19.213 -apply (simp (no_asm_simp))
  19.214 -done
  19.215 -
  19.216 -lemma drop_type [simp,TC]: "[| i:nat; l: list(A) |] ==> drop(i,l) : list(A)"
  19.217 -apply (induct_tac "i")
  19.218 -apply (simp_all (no_asm_simp) add: tl_type)
  19.219 -done
  19.220 -
  19.221 -declare drop_succ [simp del]
  19.222 -
  19.223 -
  19.224 -(** Type checking -- proved by induction, as usual **)
  19.225 -
  19.226 -lemma list_rec_type [TC]:
  19.227 -    "[| l: list(A);
  19.228 -        c: C(Nil);
  19.229 -        !!x y r. [| x:A;  y: list(A);  r: C(y) |] ==> h(x,y,r): C(Cons(x,y))
  19.230 -     |] ==> list_rec(c,h,l) : C(l)"
  19.231 -by (induct_tac "l", auto)
  19.232 -
  19.233 -(** map **)
  19.234 -
  19.235 -lemma map_type [TC]:
  19.236 -    "[| l: list(A);  !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)"
  19.237 -apply (simp add: map_list_def)
  19.238 -apply (typecheck add: list.intros list_rec_type, blast)
  19.239 -done
  19.240 -
  19.241 -lemma map_type2 [TC]: "l: list(A) ==> map(h,l) : list({h(u). u:A})"
  19.242 -apply (erule map_type)
  19.243 -apply (erule RepFunI)
  19.244 -done
  19.245 -
  19.246 -(** length **)
  19.247 -
  19.248 -lemma length_type [TC]: "l: list(A) ==> length(l) : nat"
  19.249 -by (simp add: length_list_def)
  19.250 -
  19.251 -lemma lt_length_in_nat:
  19.252 -   "[|x < length(xs); xs \<in> list(A)|] ==> x \<in> nat"
  19.253 -by (frule lt_nat_in_nat, typecheck) 
  19.254 -
  19.255 -(** app **)
  19.256 -
  19.257 -lemma app_type [TC]: "[| xs: list(A);  ys: list(A) |] ==> xs@ys : list(A)"
  19.258 -by (simp add: app_list_def)
  19.259 -
  19.260 -(** rev **)
  19.261 -
  19.262 -lemma rev_type [TC]: "xs: list(A) ==> rev(xs) : list(A)"
  19.263 -by (simp add: rev_list_def)
  19.264 -
  19.265 -
  19.266 -(** flat **)
  19.267 -
  19.268 -lemma flat_type [TC]: "ls: list(list(A)) ==> flat(ls) : list(A)"
  19.269 -by (simp add: flat_list_def)
  19.270 -
  19.271 -
  19.272 -(** set_of_list **)
  19.273 -
  19.274 -lemma set_of_list_type [TC]: "l: list(A) ==> set_of_list(l) : Pow(A)"
  19.275 -apply (unfold set_of_list_list_def)
  19.276 -apply (erule list_rec_type, auto)
  19.277 -done
  19.278 -
  19.279 -lemma set_of_list_append:
  19.280 -     "xs: list(A) ==> set_of_list (xs@ys) = set_of_list(xs) Un set_of_list(ys)"
  19.281 -apply (erule list.induct)
  19.282 -apply (simp_all (no_asm_simp) add: Un_cons)
  19.283 -done
  19.284 -
  19.285 -
  19.286 -(** list_add **)
  19.287 -
  19.288 -lemma list_add_type [TC]: "xs: list(nat) ==> list_add(xs) : nat"
  19.289 -by (simp add: list_add_list_def)
  19.290 -
  19.291 -
  19.292 -(*** theorems about map ***)
  19.293 -
  19.294 -lemma map_ident [simp]: "l: list(A) ==> map(%u. u, l) = l"
  19.295 -apply (induct_tac "l")
  19.296 -apply (simp_all (no_asm_simp))
  19.297 -done
  19.298 -
  19.299 -lemma map_compose: "l: list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)"
  19.300 -apply (induct_tac "l")
  19.301 -apply (simp_all (no_asm_simp))
  19.302 -done
  19.303 -
  19.304 -lemma map_app_distrib: "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)"
  19.305 -apply (induct_tac "xs")
  19.306 -apply (simp_all (no_asm_simp))
  19.307 -done
  19.308 -
  19.309 -lemma map_flat: "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))"
  19.310 -apply (induct_tac "ls")
  19.311 -apply (simp_all (no_asm_simp) add: map_app_distrib)
  19.312 -done
  19.313 -
  19.314 -lemma list_rec_map:
  19.315 -     "l: list(A) ==>
  19.316 -      list_rec(c, d, map(h,l)) =
  19.317 -      list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)"
  19.318 -apply (induct_tac "l")
  19.319 -apply (simp_all (no_asm_simp))
  19.320 -done
  19.321 -
  19.322 -(** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **)
  19.323 -
  19.324 -(* c : list(Collect(B,P)) ==> c : list(B) *)
  19.325 -lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD, standard]
  19.326 -
  19.327 -lemma map_list_Collect: "l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)"
  19.328 -apply (induct_tac "l")
  19.329 -apply (simp_all (no_asm_simp))
  19.330 -done
  19.331 -
  19.332 -(*** theorems about length ***)
  19.333 -
  19.334 -lemma length_map [simp]: "xs: list(A) ==> length(map(h,xs)) = length(xs)"
  19.335 -by (induct_tac "xs", simp_all)
  19.336 -
  19.337 -lemma length_app [simp]:
  19.338 -     "[| xs: list(A); ys: list(A) |]
  19.339 -      ==> length(xs@ys) = length(xs) #+ length(ys)"
  19.340 -by (induct_tac "xs", simp_all)
  19.341 -
  19.342 -lemma length_rev [simp]: "xs: list(A) ==> length(rev(xs)) = length(xs)"
  19.343 -apply (induct_tac "xs")
  19.344 -apply (simp_all (no_asm_simp) add: length_app)
  19.345 -done
  19.346 -
  19.347 -lemma length_flat:
  19.348 -     "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))"
  19.349 -apply (induct_tac "ls")
  19.350 -apply (simp_all (no_asm_simp) add: length_app)
  19.351 -done
  19.352 -
  19.353 -(** Length and drop **)
  19.354 -
  19.355 -(*Lemma for the inductive step of drop_length*)
  19.356 -lemma drop_length_Cons [rule_format]:
  19.357 -     "xs: list(A) ==>
  19.358 -           \<forall>x.  EX z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)"
  19.359 -by (erule list.induct, simp_all)
  19.360 -
  19.361 -lemma drop_length [rule_format]:
  19.362 -     "l: list(A) ==> \<forall>i \<in> length(l). (EX z zs. drop(i,l) = Cons(z,zs))"
  19.363 -apply (erule list.induct, simp_all, safe)
  19.364 -apply (erule drop_length_Cons)
  19.365 -apply (rule natE)
  19.366 -apply (erule Ord_trans [OF asm_rl length_type Ord_nat], assumption, simp_all)
  19.367 -apply (blast intro: succ_in_naturalD length_type)
  19.368 -done
  19.369 -
  19.370 -
  19.371 -(*** theorems about app ***)
  19.372 -
  19.373 -lemma app_right_Nil [simp]: "xs: list(A) ==> xs@Nil=xs"
  19.374 -by (erule list.induct, simp_all)
  19.375 -
  19.376 -lemma app_assoc: "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)"
  19.377 -by (induct_tac "xs", simp_all)
  19.378 -
  19.379 -lemma flat_app_distrib: "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)"
  19.380 -apply (induct_tac "ls")
  19.381 -apply (simp_all (no_asm_simp) add: app_assoc)
  19.382 -done
  19.383 -
  19.384 -(*** theorems about rev ***)
  19.385 -
  19.386 -lemma rev_map_distrib: "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))"
  19.387 -apply (induct_tac "l")
  19.388 -apply (simp_all (no_asm_simp) add: map_app_distrib)
  19.389 -done
  19.390 -
  19.391 -(*Simplifier needs the premises as assumptions because rewriting will not
  19.392 -  instantiate the variable ?A in the rules' typing conditions; note that
  19.393 -  rev_type does not instantiate ?A.  Only the premises do.
  19.394 -*)
  19.395 -lemma rev_app_distrib:
  19.396 -     "[| xs: list(A);  ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"
  19.397 -apply (erule list.induct)
  19.398 -apply (simp_all add: app_assoc)
  19.399 -done
  19.400 -
  19.401 -lemma rev_rev_ident [simp]: "l: list(A) ==> rev(rev(l))=l"
  19.402 -apply (induct_tac "l")
  19.403 -apply (simp_all (no_asm_simp) add: rev_app_distrib)
  19.404 -done
  19.405 -
  19.406 -lemma rev_flat: "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))"
  19.407 -apply (induct_tac "ls")
  19.408 -apply (simp_all add: map_app_distrib flat_app_distrib rev_app_distrib)
  19.409 -done
  19.410 -
  19.411 -
  19.412 -(*** theorems about list_add ***)
  19.413 -
  19.414 -lemma list_add_app:
  19.415 -     "[| xs: list(nat);  ys: list(nat) |]
  19.416 -      ==> list_add(xs@ys) = list_add(ys) #+ list_add(xs)"
  19.417 -apply (induct_tac "xs", simp_all)
  19.418 -done
  19.419 -
  19.420 -lemma list_add_rev: "l: list(nat) ==> list_add(rev(l)) = list_add(l)"
  19.421 -apply (induct_tac "l")
  19.422 -apply (simp_all (no_asm_simp) add: list_add_app)
  19.423 -done
  19.424 -
  19.425 -lemma list_add_flat:
  19.426 -     "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))"
  19.427 -apply (induct_tac "ls")
  19.428 -apply (simp_all (no_asm_simp) add: list_add_app)
  19.429 -done
  19.430 -
  19.431 -(** New induction rules **)
  19.432 -
  19.433 -lemma list_append_induct [case_names Nil snoc, consumes 1]:
  19.434 -    "[| l: list(A);
  19.435 -        P(Nil);
  19.436 -        !!x y. [| x: A;  y: list(A);  P(y) |] ==> P(y @ [x])
  19.437 -     |] ==> P(l)"
  19.438 -apply (subgoal_tac "P(rev(rev(l)))", simp)
  19.439 -apply (erule rev_type [THEN list.induct], simp_all)
  19.440 -done
  19.441 -
  19.442 -lemma list_complete_induct_lemma [rule_format]:
  19.443 - assumes ih: 
  19.444 -    "\<And>l. [| l \<in> list(A); 
  19.445 -             \<forall>l' \<in> list(A). length(l') < length(l) --> P(l')|] 
  19.446 -          ==> P(l)"
  19.447 -  shows "n \<in> nat ==> \<forall>l \<in> list(A). length(l) < n --> P(l)"
  19.448 -apply (induct_tac n, simp)
  19.449 -apply (blast intro: ih elim!: leE) 
  19.450 -done
  19.451 -
  19.452 -theorem list_complete_induct:
  19.453 -      "[| l \<in> list(A); 
  19.454 -          \<And>l. [| l \<in> list(A); 
  19.455 -                  \<forall>l' \<in> list(A). length(l') < length(l) --> P(l')|] 
  19.456 -               ==> P(l)
  19.457 -       |] ==> P(l)"
  19.458 -apply (rule list_complete_induct_lemma [of A]) 
  19.459 -   prefer 4 apply (rule le_refl, simp) 
  19.460 -  apply blast 
  19.461 - apply simp 
  19.462 -apply assumption
  19.463 -done
  19.464 -
  19.465 -
  19.466 -(*** Thanks to Sidi Ehmety for these results about min, take, etc. ***)
  19.467 -
  19.468 -(** min FIXME: replace by Int! **)
  19.469 -(* Min theorems are also true for i, j ordinals *)
  19.470 -lemma min_sym: "[| i:nat; j:nat |] ==> min(i,j)=min(j,i)"
  19.471 -apply (unfold min_def)
  19.472 -apply (auto dest!: not_lt_imp_le dest: lt_not_sym intro: le_anti_sym)
  19.473 -done
  19.474 -
  19.475 -lemma min_type [simp,TC]: "[| i:nat; j:nat |] ==> min(i,j):nat"
  19.476 -by (unfold min_def, auto)
  19.477 -
  19.478 -lemma min_0 [simp]: "i:nat ==> min(0,i) = 0"
  19.479 -apply (unfold min_def)
  19.480 -apply (auto dest: not_lt_imp_le)
  19.481 -done
  19.482 -
  19.483 -lemma min_02 [simp]: "i:nat ==> min(i, 0) = 0"
  19.484 -apply (unfold min_def)
  19.485 -apply (auto dest: not_lt_imp_le)
  19.486 -done
  19.487 -
  19.488 -lemma lt_min_iff: "[| i:nat; j:nat; k:nat |] ==> i<min(j,k) <-> i<j & i<k"
  19.489 -apply (unfold min_def)
  19.490 -apply (auto dest!: not_lt_imp_le intro: lt_trans2 lt_trans)
  19.491 -done
  19.492 -
  19.493 -lemma min_succ_succ [simp]:
  19.494 -     "[| i:nat; j:nat |] ==>  min(succ(i), succ(j))= succ(min(i, j))"
  19.495 -apply (unfold min_def, auto)
  19.496 -done
  19.497 -
  19.498 -(*** more theorems about lists ***)
  19.499 -
  19.500 -(** filter **)
  19.501 -
  19.502 -lemma filter_append [simp]:
  19.503 -     "xs:list(A) ==> filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)"
  19.504 -by (induct_tac "xs", auto)
  19.505 -
  19.506 -lemma filter_type [simp,TC]: "xs:list(A) ==> filter(P, xs):list(A)"
  19.507 -by (induct_tac "xs", auto)
  19.508 -
  19.509 -lemma length_filter: "xs:list(A) ==> length(filter(P, xs)) le length(xs)"
  19.510 -apply (induct_tac "xs", auto)
  19.511 -apply (rule_tac j = "length (l) " in le_trans)
  19.512 -apply (auto simp add: le_iff)
  19.513 -done
  19.514 -
  19.515 -lemma filter_is_subset: "xs:list(A) ==> set_of_list(filter(P,xs)) <= set_of_list(xs)"
  19.516 -by (induct_tac "xs", auto)
  19.517 -
  19.518 -lemma filter_False [simp]: "xs:list(A) ==> filter(%p. False, xs) = Nil"
  19.519 -by (induct_tac "xs", auto)
  19.520 -
  19.521 -lemma filter_True [simp]: "xs:list(A) ==> filter(%p. True, xs) = xs"
  19.522 -by (induct_tac "xs", auto)
  19.523 -
  19.524 -(** length **)
  19.525 -
  19.526 -lemma length_is_0_iff [simp]: "xs:list(A) ==> length(xs)=0 <-> xs=Nil"
  19.527 -by (erule list.induct, auto)
  19.528 -
  19.529 -lemma length_is_0_iff2 [simp]: "xs:list(A) ==> 0 = length(xs) <-> xs=Nil"
  19.530 -by (erule list.induct, auto)
  19.531 -
  19.532 -lemma length_tl [simp]: "xs:list(A) ==> length(tl(xs)) = length(xs) #- 1"
  19.533 -by (erule list.induct, auto)
  19.534 -
  19.535 -lemma length_greater_0_iff: "xs:list(A) ==> 0<length(xs) <-> xs ~= Nil"
  19.536 -by (erule list.induct, auto)
  19.537 -
  19.538 -lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) <-> (EX y ys. xs=Cons(y, ys) & length(ys)=n)"
  19.539 -by (erule list.induct, auto)
  19.540 -
  19.541 -(** more theorems about append **)
  19.542 -
  19.543 -lemma append_is_Nil_iff [simp]:
  19.544 -     "xs:list(A) ==> (xs@ys = Nil) <-> (xs=Nil & ys = Nil)"
  19.545 -by (erule list.induct, auto)
  19.546 -
  19.547 -lemma append_is_Nil_iff2 [simp]:
  19.548 -     "xs:list(A) ==> (Nil = xs@ys) <-> (xs=Nil & ys = Nil)"
  19.549 -by (erule list.induct, auto)
  19.550 -
  19.551 -lemma append_left_is_self_iff [simp]:
  19.552 -     "xs:list(A) ==> (xs@ys = xs) <-> (ys = Nil)"
  19.553 -by (erule list.induct, auto)
  19.554 -
  19.555 -lemma append_left_is_self_iff2 [simp]:
  19.556 -     "xs:list(A) ==> (xs = xs@ys) <-> (ys = Nil)"
  19.557 -by (erule list.induct, auto)
  19.558 -
  19.559 -(*TOO SLOW as a default simprule!*)
  19.560 -lemma append_left_is_Nil_iff [rule_format]:
  19.561 -     "[| xs:list(A); ys:list(A); zs:list(A) |] ==>
  19.562 -   length(ys)=length(zs) --> (xs@ys=zs <-> (xs=Nil & ys=zs))"
  19.563 -apply (erule list.induct)
  19.564 -apply (auto simp add: length_app)
  19.565 -done
  19.566 -
  19.567 -(*TOO SLOW as a default simprule!*)
  19.568 -lemma append_left_is_Nil_iff2 [rule_format]:
  19.569 -     "[| xs:list(A); ys:list(A); zs:list(A) |] ==>
  19.570 -   length(ys)=length(zs) --> (zs=ys@xs <-> (xs=Nil & ys=zs))"
  19.571 -apply (erule list.induct)
  19.572 -apply (auto simp add: length_app)
  19.573 -done
  19.574 -
  19.575 -lemma append_eq_append_iff [rule_format,simp]:
  19.576 -     "xs:list(A) ==> \<forall>ys \<in> list(A).
  19.577 -      length(xs)=length(ys) --> (xs@us = ys@vs) <-> (xs=ys & us=vs)"
  19.578 -apply (erule list.induct)
  19.579 -apply (simp (no_asm_simp))
  19.580 -apply clarify
  19.581 -apply (erule_tac a = ys in list.cases, auto)
  19.582 -done
  19.583 -
  19.584 -lemma append_eq_append [rule_format]:
  19.585 -  "xs:list(A) ==>
  19.586 -   \<forall>ys \<in> list(A). \<forall>us \<in> list(A). \<forall>vs \<in> list(A).
  19.587 -   length(us) = length(vs) --> (xs@us = ys@vs) --> (xs=ys & us=vs)"
  19.588 -apply (induct_tac "xs")
  19.589 -apply (force simp add: length_app, clarify)
  19.590 -apply (erule_tac a = ys in list.cases, simp)
  19.591 -apply (subgoal_tac "Cons (a, l) @ us =vs")
  19.592 - apply (drule rev_iffD1 [OF _ append_left_is_Nil_iff], simp_all, blast)
  19.593 -done
  19.594 -
  19.595 -lemma append_eq_append_iff2 [simp]:
  19.596 - "[| xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs) |]
  19.597 -  ==>  xs@us = ys@vs <-> (xs=ys & us=vs)"
  19.598 -apply (rule iffI)
  19.599 -apply (rule append_eq_append, auto)
  19.600 -done
  19.601 -
  19.602 -lemma append_self_iff [simp]:
  19.603 -     "[| xs:list(A); ys:list(A); zs:list(A) |] ==> xs@ys=xs@zs <-> ys=zs"
  19.604 -by simp
  19.605 -
  19.606 -lemma append_self_iff2 [simp]:
  19.607 -     "[| xs:list(A); ys:list(A); zs:list(A) |] ==> ys@xs=zs@xs <-> ys=zs"
  19.608 -by simp
  19.609 -
  19.610 -(* Can also be proved from append_eq_append_iff2,
  19.611 -but the proof requires two more hypotheses: x:A and y:A *)
  19.612 -lemma append1_eq_iff [rule_format,simp]:
  19.613 -     "xs:list(A) ==> \<forall>ys \<in> list(A). xs@[x] = ys@[y] <-> (xs = ys & x=y)"
  19.614 -apply (erule list.induct)  
  19.615 - apply clarify 
  19.616 - apply (erule list.cases)
  19.617 - apply simp_all
  19.618 -txt{*Inductive step*}  
  19.619 -apply clarify 
  19.620 -apply (erule_tac a=ys in list.cases, simp_all)  
  19.621 -done
  19.622 -
  19.623 -
  19.624 -lemma append_right_is_self_iff [simp]:
  19.625 -     "[| xs:list(A); ys:list(A) |] ==> (xs@ys = ys) <-> (xs=Nil)"
  19.626 -by (simp (no_asm_simp) add: append_left_is_Nil_iff)
  19.627 -
  19.628 -lemma append_right_is_self_iff2 [simp]:
  19.629 -     "[| xs:list(A); ys:list(A) |] ==> (ys = xs@ys) <-> (xs=Nil)"
  19.630 -apply (rule iffI)
  19.631 -apply (drule sym, auto) 
  19.632 -done
  19.633 -
  19.634 -lemma hd_append [rule_format,simp]:
  19.635 -     "xs:list(A) ==> xs ~= Nil --> hd(xs @ ys) = hd(xs)"
  19.636 -by (induct_tac "xs", auto)
  19.637 -
  19.638 -lemma tl_append [rule_format,simp]:
  19.639 -     "xs:list(A) ==> xs~=Nil --> tl(xs @ ys) = tl(xs)@ys"
  19.640 -by (induct_tac "xs", auto)
  19.641 -
  19.642 -(** rev **)
  19.643 -lemma rev_is_Nil_iff [simp]: "xs:list(A) ==> (rev(xs) = Nil <-> xs = Nil)"
  19.644 -by (erule list.induct, auto)
  19.645 -
  19.646 -lemma Nil_is_rev_iff [simp]: "xs:list(A) ==> (Nil = rev(xs) <-> xs = Nil)"
  19.647 -by (erule list.induct, auto)
  19.648 -
  19.649 -lemma rev_is_rev_iff [rule_format,simp]:
  19.650 -     "xs:list(A) ==> \<forall>ys \<in> list(A). rev(xs)=rev(ys) <-> xs=ys"
  19.651 -apply (erule list.induct, force, clarify)
  19.652 -apply (erule_tac a = ys in list.cases, auto)
  19.653 -done
  19.654 -
  19.655 -lemma rev_list_elim [rule_format]:
  19.656 -     "xs:list(A) ==>
  19.657 -      (xs=Nil --> P) --> (\<forall>ys \<in> list(A). \<forall>y \<in> A. xs =ys@[y] -->P)-->P"
  19.658 -by (erule list_append_induct, auto)
  19.659 -
  19.660 -
  19.661 -(** more theorems about drop **)
  19.662 -
  19.663 -lemma length_drop [rule_format,simp]:
  19.664 -     "n:nat ==> \<forall>xs \<in> list(A). length(drop(n, xs)) = length(xs) #- n"
  19.665 -apply (erule nat_induct)
  19.666 -apply (auto elim: list.cases)
  19.667 -done
  19.668 -
  19.669 -lemma drop_all [rule_format,simp]:
  19.670 -     "n:nat ==> \<forall>xs \<in> list(A). length(xs) le n --> drop(n, xs)=Nil"
  19.671 -apply (erule nat_induct)
  19.672 -apply (auto elim: list.cases)
  19.673 -done
  19.674 -
  19.675 -lemma drop_append [rule_format]:
  19.676 -     "n:nat ==> 
  19.677 -      \<forall>xs \<in> list(A). drop(n, xs@ys) = drop(n,xs) @ drop(n #- length(xs), ys)"
  19.678 -apply (induct_tac "n")
  19.679 -apply (auto elim: list.cases)
  19.680 -done
  19.681 -
  19.682 -lemma drop_drop:
  19.683 -    "m:nat ==> \<forall>xs \<in> list(A). \<forall>n \<in> nat. drop(n, drop(m, xs))=drop(n #+ m, xs)"
  19.684 -apply (induct_tac "m")
  19.685 -apply (auto elim: list.cases)
  19.686 -done
  19.687 -
  19.688 -(** take **)
  19.689 -
  19.690 -lemma take_0 [simp]: "xs:list(A) ==> take(0, xs) =  Nil"
  19.691 -apply (unfold take_def)
  19.692 -apply (erule list.induct, auto)
  19.693 -done
  19.694 -
  19.695 -lemma take_succ_Cons [simp]:
  19.696 -    "n:nat ==> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))"
  19.697 -by (simp add: take_def)
  19.698 -
  19.699 -(* Needed for proving take_all *)
  19.700 -lemma take_Nil [simp]: "n:nat ==> take(n, Nil) = Nil"
  19.701 -by (unfold take_def, auto)
  19.702 -
  19.703 -lemma take_all [rule_format,simp]:
  19.704 -     "n:nat ==> \<forall>xs \<in> list(A). length(xs) le n  --> take(n, xs) = xs"
  19.705 -apply (erule nat_induct)
  19.706 -apply (auto elim: list.cases) 
  19.707 -done
  19.708 -
  19.709 -lemma take_type [rule_format,simp,TC]:
  19.710 -     "xs:list(A) ==> \<forall>n \<in> nat. take(n, xs):list(A)"
  19.711 -apply (erule list.induct, simp, clarify) 
  19.712 -apply (erule natE, auto)
  19.713 -done
  19.714 -
  19.715 -lemma take_append [rule_format,simp]:
  19.716 - "xs:list(A) ==>
  19.717 -  \<forall>ys \<in> list(A). \<forall>n \<in> nat. take(n, xs @ ys) =
  19.718 -                            take(n, xs) @ take(n #- length(xs), ys)"
  19.719 -apply (erule list.induct, simp, clarify) 
  19.720 -apply (erule natE, auto)
  19.721 -done
  19.722 -
  19.723 -lemma take_take [rule_format]:
  19.724 -   "m : nat ==>
  19.725 -    \<forall>xs \<in> list(A). \<forall>n \<in> nat. take(n, take(m,xs))= take(min(n, m), xs)"
  19.726 -apply (induct_tac "m", auto)
  19.727 -apply (erule_tac a = xs in list.cases)
  19.728 -apply (auto simp add: take_Nil)
  19.729 -apply (erule_tac n=n in natE)
  19.730 -apply (auto intro: take_0 take_type)
  19.731 -done
  19.732 -
  19.733 -(** nth **)
  19.734 -
  19.735 -lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a"
  19.736 -by (simp add: nth_def) 
  19.737 -
  19.738 -lemma nth_Cons [simp]: "n:nat ==> nth(succ(n), Cons(a,l)) = nth(n,l)"
  19.739 -by (simp add: nth_def) 
  19.740 -
  19.741 -lemma nth_empty [simp]: "nth(n, Nil) = 0"
  19.742 -by (simp add: nth_def) 
  19.743 -
  19.744 -lemma nth_type [rule_format,simp,TC]:
  19.745 -     "xs:list(A) ==> \<forall>n. n < length(xs) --> nth(n,xs) : A"
  19.746 -apply (erule list.induct, simp, clarify)
  19.747 -apply (subgoal_tac "n \<in> nat")  
  19.748 - apply (erule natE, auto dest!: le_in_nat)
  19.749 -done
  19.750 -
  19.751 -lemma nth_eq_0 [rule_format]:
  19.752 -     "xs:list(A) ==> \<forall>n \<in> nat. length(xs) le n --> nth(n,xs) = 0"
  19.753 -apply (erule list.induct, simp, clarify) 
  19.754 -apply (erule natE, auto)
  19.755 -done
  19.756 -
  19.757 -lemma nth_append [rule_format]:
  19.758 -  "xs:list(A) ==> 
  19.759 -   \<forall>n \<in> nat. nth(n, xs @ ys) = (if n < length(xs) then nth(n,xs)
  19.760 -                                else nth(n #- length(xs), ys))"
  19.761 -apply (induct_tac "xs", simp, clarify) 
  19.762 -apply (erule natE, auto)
  19.763 -done
  19.764 -
  19.765 -lemma set_of_list_conv_nth:
  19.766 -    "xs:list(A)
  19.767 -     ==> set_of_list(xs) = {x:A. EX i:nat. i<length(xs) & x = nth(i,xs)}"
  19.768 -apply (induct_tac "xs", simp_all)
  19.769 -apply (rule equalityI, auto)
  19.770 -apply (rule_tac x = 0 in bexI, auto)
  19.771 -apply (erule natE, auto)
  19.772 -done
  19.773 -
  19.774 -(* Other theorems about lists *)
  19.775 -
  19.776 -lemma nth_take_lemma [rule_format]:
  19.777 - "k:nat ==>
  19.778 -  \<forall>xs \<in> list(A). (\<forall>ys \<in> list(A). k le length(xs) --> k le length(ys) -->
  19.779 -      (\<forall>i \<in> nat. i<k --> nth(i,xs) = nth(i,ys))--> take(k,xs) = take(k,ys))"
  19.780 -apply (induct_tac "k")
  19.781 -apply (simp_all (no_asm_simp) add: lt_succ_eq_0_disj all_conj_distrib)
  19.782 -apply clarify
  19.783 -(*Both lists are non-empty*)
  19.784 -apply (erule_tac a=xs in list.cases, simp) 
  19.785 -apply (erule_tac a=ys in list.cases, clarify) 
  19.786 -apply (simp (no_asm_use) )
  19.787 -apply clarify
  19.788 -apply (simp (no_asm_simp))
  19.789 -apply (rule conjI, force)
  19.790 -apply (rename_tac y ys z zs) 
  19.791 -apply (drule_tac x = zs and x1 = ys in bspec [THEN bspec], auto)   
  19.792 -done
  19.793 -
  19.794 -lemma nth_equalityI [rule_format]:
  19.795 -     "[| xs:list(A); ys:list(A); length(xs) = length(ys);
  19.796 -         \<forall>i \<in> nat. i < length(xs) --> nth(i,xs) = nth(i,ys) |]
  19.797 -      ==> xs = ys"
  19.798 -apply (subgoal_tac "length (xs) le length (ys) ")
  19.799 -apply (cut_tac k="length(xs)" and xs=xs and ys=ys in nth_take_lemma) 
  19.800 -apply (simp_all add: take_all)
  19.801 -done
  19.802 -
  19.803 -(*The famous take-lemma*)
  19.804 -
  19.805 -lemma take_equalityI [rule_format]:
  19.806 -    "[| xs:list(A); ys:list(A); (\<forall>i \<in> nat. take(i, xs) = take(i,ys)) |] 
  19.807 -     ==> xs = ys"
  19.808 -apply (case_tac "length (xs) le length (ys) ")
  19.809 -apply (drule_tac x = "length (ys) " in bspec)
  19.810 -apply (drule_tac [3] not_lt_imp_le)
  19.811 -apply (subgoal_tac [5] "length (ys) le length (xs) ")
  19.812 -apply (rule_tac [6] j = "succ (length (ys))" in le_trans)
  19.813 -apply (rule_tac [6] leI)
  19.814 -apply (drule_tac [5] x = "length (xs) " in bspec)
  19.815 -apply (simp_all add: take_all)
  19.816 -done
  19.817 -
  19.818 -lemma nth_drop [rule_format]:
  19.819 -  "n:nat ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)"
  19.820 -apply (induct_tac "n", simp_all, clarify)
  19.821 -apply (erule list.cases, auto)  
  19.822 -done
  19.823 -
  19.824 -lemma take_succ [rule_format]:
  19.825 -  "xs\<in>list(A) 
  19.826 -   ==> \<forall>i. i < length(xs) --> take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]"
  19.827 -apply (induct_tac "xs", auto)
  19.828 -apply (subgoal_tac "i\<in>nat") 
  19.829 -apply (erule natE)
  19.830 -apply (auto simp add: le_in_nat) 
  19.831 -done
  19.832 -
  19.833 -lemma take_add [rule_format]:
  19.834 -     "[|xs\<in>list(A); j\<in>nat|] 
  19.835 -      ==> \<forall>i\<in>nat. take(i #+ j, xs) = take(i,xs) @ take(j, drop(i,xs))"
  19.836 -apply (induct_tac "xs", simp_all, clarify)
  19.837 -apply (erule_tac n = i in natE, simp_all)
  19.838 -done
  19.839 -
  19.840 -lemma length_take:
  19.841 -     "l\<in>list(A) ==> \<forall>n\<in>nat. length(take(n,l)) = min(n, length(l))"
  19.842 -apply (induct_tac "l", safe, simp_all)
  19.843 -apply (erule natE, simp_all)
  19.844 -done
  19.845 -
  19.846 -subsection{*The function zip*}
  19.847 -
  19.848 -text{*Crafty definition to eliminate a type argument*}
  19.849 -
  19.850 -consts
  19.851 -  zip_aux        :: "[i,i]=>i"
  19.852 -
  19.853 -primrec (*explicit lambda is required because both arguments of "un" vary*)
  19.854 -  "zip_aux(B,[]) =
  19.855 -     (\<lambda>ys \<in> list(B). list_case([], %y l. [], ys))"
  19.856 -
  19.857 -  "zip_aux(B,Cons(x,l)) =
  19.858 -     (\<lambda>ys \<in> list(B).
  19.859 -        list_case(Nil, %y zs. Cons(<x,y>, zip_aux(B,l)`zs), ys))"
  19.860 -
  19.861 -definition
  19.862 -  zip :: "[i, i]=>i"  where
  19.863 -   "zip(xs, ys) == zip_aux(set_of_list(ys),xs)`ys"
  19.864 -
  19.865 -
  19.866 -(* zip equations *)
  19.867 -
  19.868 -lemma list_on_set_of_list: "xs \<in> list(A) ==> xs \<in> list(set_of_list(xs))"
  19.869 -apply (induct_tac xs, simp_all) 
  19.870 -apply (blast intro: list_mono [THEN subsetD]) 
  19.871 -done
  19.872 -
  19.873 -lemma zip_Nil [simp]: "ys:list(A) ==> zip(Nil, ys)=Nil"
  19.874 -apply (simp add: zip_def list_on_set_of_list [of _ A])
  19.875 -apply (erule list.cases, simp_all) 
  19.876 -done
  19.877 -
  19.878 -lemma zip_Nil2 [simp]: "xs:list(A) ==> zip(xs, Nil)=Nil"
  19.879 -apply (simp add: zip_def list_on_set_of_list [of _ A])
  19.880 -apply (erule list.cases, simp_all) 
  19.881 -done
  19.882 -
  19.883 -lemma zip_aux_unique [rule_format]:
  19.884 -     "[|B<=C;  xs \<in> list(A)|] 
  19.885 -      ==> \<forall>ys \<in> list(B). zip_aux(C,xs) ` ys = zip_aux(B,xs) ` ys"
  19.886 -apply (induct_tac xs) 
  19.887 - apply simp_all 
  19.888 - apply (blast intro: list_mono [THEN subsetD], clarify) 
  19.889 -apply (erule_tac a=ys in list.cases, auto) 
  19.890 -apply (blast intro: list_mono [THEN subsetD]) 
  19.891 -done
  19.892 -
  19.893 -lemma zip_Cons_Cons [simp]:
  19.894 -     "[| xs:list(A); ys:list(B); x:A; y:B |] ==>
  19.895 -      zip(Cons(x,xs), Cons(y, ys)) = Cons(<x,y>, zip(xs, ys))"
  19.896 -apply (simp add: zip_def, auto) 
  19.897 -apply (rule zip_aux_unique, auto) 
  19.898 -apply (simp add: list_on_set_of_list [of _ B])
  19.899 -apply (blast intro: list_on_set_of_list list_mono [THEN subsetD]) 
  19.900 -done
  19.901 -
  19.902 -lemma zip_type [rule_format,simp,TC]:
  19.903 -     "xs:list(A) ==> \<forall>ys \<in> list(B). zip(xs, ys):list(A*B)"
  19.904 -apply (induct_tac "xs")
  19.905 -apply (simp (no_asm))
  19.906 -apply clarify
  19.907 -apply (erule_tac a = ys in list.cases, auto)
  19.908 -done
  19.909 -
  19.910 -(* zip length *)
  19.911 -lemma length_zip [rule_format,simp]:
  19.912 -     "xs:list(A) ==> \<forall>ys \<in> list(B). length(zip(xs,ys)) =
  19.913 -                                     min(length(xs), length(ys))"
  19.914 -apply (unfold min_def)
  19.915 -apply (induct_tac "xs", simp_all, clarify) 
  19.916 -apply (erule_tac a = ys in list.cases, auto)
  19.917 -done
  19.918 -
  19.919 -lemma zip_append1 [rule_format]:
  19.920 - "[| ys:list(A); zs:list(B) |] ==>
  19.921 -  \<forall>xs \<in> list(A). zip(xs @ ys, zs) = 
  19.922 -                 zip(xs, take(length(xs), zs)) @ zip(ys, drop(length(xs),zs))"
  19.923 -apply (induct_tac "zs", force, clarify) 
  19.924 -apply (erule_tac a = xs in list.cases, simp_all) 
  19.925 -done
  19.926 -
  19.927 -lemma zip_append2 [rule_format]:
  19.928 - "[| xs:list(A); zs:list(B) |] ==> \<forall>ys \<in> list(B). zip(xs, ys@zs) =
  19.929 -       zip(take(length(ys), xs), ys) @ zip(drop(length(ys), xs), zs)"
  19.930 -apply (induct_tac "xs", force, clarify) 
  19.931 -apply (erule_tac a = ys in list.cases, auto)
  19.932 -done
  19.933 -
  19.934 -lemma zip_append [simp]:
  19.935 - "[| length(xs) = length(us); length(ys) = length(vs);
  19.936 -     xs:list(A); us:list(B); ys:list(A); vs:list(B) |] 
  19.937 -  ==> zip(xs@ys,us@vs) = zip(xs, us) @ zip(ys, vs)"
  19.938 -by (simp (no_asm_simp) add: zip_append1 drop_append diff_self_eq_0)
  19.939 -
  19.940 -
  19.941 -lemma zip_rev [rule_format,simp]:
  19.942 - "ys:list(B) ==> \<forall>xs \<in> list(A).
  19.943 -    length(xs) = length(ys) --> zip(rev(xs), rev(ys)) = rev(zip(xs, ys))"
  19.944 -apply (induct_tac "ys", force, clarify) 
  19.945 -apply (erule_tac a = xs in list.cases)
  19.946 -apply (auto simp add: length_rev)
  19.947 -done
  19.948 -
  19.949 -lemma nth_zip [rule_format,simp]:
  19.950 -   "ys:list(B) ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A).
  19.951 -                    i < length(xs) --> i < length(ys) -->
  19.952 -                    nth(i,zip(xs, ys)) = <nth(i,xs),nth(i, ys)>"
  19.953 -apply (induct_tac "ys", force, clarify) 
  19.954 -apply (erule_tac a = xs in list.cases, simp) 
  19.955 -apply (auto elim: natE)
  19.956 -done
  19.957 -
  19.958 -lemma set_of_list_zip [rule_format]:
  19.959 -     "[| xs:list(A); ys:list(B); i:nat |]
  19.960 -      ==> set_of_list(zip(xs, ys)) =
  19.961 -          {<x, y>:A*B. EX i:nat. i < min(length(xs), length(ys))
  19.962 -          & x = nth(i, xs) & y = nth(i, ys)}"
  19.963 -by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth)
  19.964 -
  19.965 -(** list_update **)
  19.966 -
  19.967 -lemma list_update_Nil [simp]: "i:nat ==>list_update(Nil, i, v) = Nil"
  19.968 -by (unfold list_update_def, auto)
  19.969 -
  19.970 -lemma list_update_Cons_0 [simp]: "list_update(Cons(x, xs), 0, v)= Cons(v, xs)"
  19.971 -by (unfold list_update_def, auto)
  19.972 -
  19.973 -lemma list_update_Cons_succ [simp]:
  19.974 -  "n:nat ==>
  19.975 -    list_update(Cons(x, xs), succ(n), v)= Cons(x, list_update(xs, n, v))"
  19.976 -apply (unfold list_update_def, auto)
  19.977 -done
  19.978 -
  19.979 -lemma list_update_type [rule_format,simp,TC]:
  19.980 -     "[| xs:list(A); v:A |] ==> \<forall>n \<in> nat. list_update(xs, n, v):list(A)"
  19.981 -apply (induct_tac "xs")
  19.982 -apply (simp (no_asm))
  19.983 -apply clarify
  19.984 -apply (erule natE, auto)
  19.985 -done
  19.986 -
  19.987 -lemma length_list_update [rule_format,simp]:
  19.988 -     "xs:list(A) ==> \<forall>i \<in> nat. length(list_update(xs, i, v))=length(xs)"
  19.989 -apply (induct_tac "xs")
  19.990 -apply (simp (no_asm))
  19.991 -apply clarify
  19.992 -apply (erule natE, auto)
  19.993 -done
  19.994 -
  19.995 -lemma nth_list_update [rule_format]:
  19.996 -     "[| xs:list(A) |] ==> \<forall>i \<in> nat. \<forall>j \<in> nat. i < length(xs)  -->
  19.997 -         nth(j, list_update(xs, i, x)) = (if i=j then x else nth(j, xs))"
  19.998 -apply (induct_tac "xs")
  19.999 - apply simp_all
 19.1000 -apply clarify
 19.1001 -apply (rename_tac i j) 
 19.1002 -apply (erule_tac n=i in natE) 
 19.1003 -apply (erule_tac [2] n=j in natE)
 19.1004 -apply (erule_tac n=j in natE, simp_all, force) 
 19.1005 -done
 19.1006 -
 19.1007 -lemma nth_list_update_eq [simp]:
 19.1008 -     "[| i < length(xs); xs:list(A) |] ==> nth(i, list_update(xs, i,x)) = x"
 19.1009 -by (simp (no_asm_simp) add: lt_length_in_nat nth_list_update)
 19.1010 -
 19.1011 -
 19.1012 -lemma nth_list_update_neq [rule_format,simp]:
 19.1013 -  "xs:list(A) ==> 
 19.1014 -     \<forall>i \<in> nat. \<forall>j \<in> nat. i ~= j --> nth(j, list_update(xs,i,x)) = nth(j,xs)"
 19.1015 -apply (induct_tac "xs")
 19.1016 - apply (simp (no_asm))
 19.1017 -apply clarify
 19.1018 -apply (erule natE)
 19.1019 -apply (erule_tac [2] natE, simp_all) 
 19.1020 -apply (erule natE, simp_all)
 19.1021 -done
 19.1022 -
 19.1023 -lemma list_update_overwrite [rule_format,simp]:
 19.1024 -     "xs:list(A) ==> \<forall>i \<in> nat. i < length(xs)
 19.1025 -   --> list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)"
 19.1026 -apply (induct_tac "xs")
 19.1027 - apply (simp (no_asm))
 19.1028 -apply clarify
 19.1029 -apply (erule natE, auto)
 19.1030 -done
 19.1031 -
 19.1032 -lemma list_update_same_conv [rule_format]:
 19.1033 -     "xs:list(A) ==> 
 19.1034 -      \<forall>i \<in> nat. i < length(xs) --> 
 19.1035 -                 (list_update(xs, i, x) = xs) <-> (nth(i, xs) = x)"
 19.1036 -apply (induct_tac "xs")
 19.1037 - apply (simp (no_asm))
 19.1038 -apply clarify
 19.1039 -apply (erule natE, auto)
 19.1040 -done
 19.1041 -
 19.1042 -lemma update_zip [rule_format]:
 19.1043 -     "ys:list(B) ==> 
 19.1044 -      \<forall>i \<in> nat. \<forall>xy \<in> A*B. \<forall>xs \<in> list(A).
 19.1045 -        length(xs) = length(ys) -->
 19.1046 -        list_update(zip(xs, ys), i, xy) = zip(list_update(xs, i, fst(xy)),
 19.1047 -                                              list_update(ys, i, snd(xy)))"
 19.1048 -apply (induct_tac "ys")
 19.1049 - apply auto
 19.1050 -apply (erule_tac a = xs in list.cases)
 19.1051 -apply (auto elim: natE)
 19.1052 -done
 19.1053 -
 19.1054 -lemma set_update_subset_cons [rule_format]:
 19.1055 -  "xs:list(A) ==> 
 19.1056 -   \<forall>i \<in> nat. set_of_list(list_update(xs, i, x)) <= cons(x, set_of_list(xs))"
 19.1057 -apply (induct_tac "xs")
 19.1058 - apply simp
 19.1059 -apply (rule ballI)
 19.1060 -apply (erule natE, simp_all, auto)
 19.1061 -done
 19.1062 -
 19.1063 -lemma set_of_list_update_subsetI:
 19.1064 -     "[| set_of_list(xs) <= A; xs:list(A); x:A; i:nat|]
 19.1065 -   ==> set_of_list(list_update(xs, i,x)) <= A"
 19.1066 -apply (rule subset_trans)
 19.1067 -apply (rule set_update_subset_cons, auto)
 19.1068 -done
 19.1069 -
 19.1070 -(** upt **)
 19.1071 -
 19.1072 -lemma upt_rec:
 19.1073 -     "j:nat ==> upt(i,j) = (if i<j then Cons(i, upt(succ(i), j)) else Nil)"
 19.1074 -apply (induct_tac "j", auto)
 19.1075 -apply (drule not_lt_imp_le)
 19.1076 -apply (auto simp: lt_Ord intro: le_anti_sym)
 19.1077 -done
 19.1078 -
 19.1079 -lemma upt_conv_Nil [simp]: "[| j le i; j:nat |] ==> upt(i,j) = Nil"
 19.1080 -apply (subst upt_rec, auto)
 19.1081 -apply (auto simp add: le_iff)
 19.1082 -apply (drule lt_asym [THEN notE], auto)
 19.1083 -done
 19.1084 -
 19.1085 -(*Only needed if upt_Suc is deleted from the simpset*)
 19.1086 -lemma upt_succ_append:
 19.1087 -     "[| i le j; j:nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]"
 19.1088 -by simp
 19.1089 -
 19.1090 -lemma upt_conv_Cons:
 19.1091 -     "[| i<j; j:nat |]  ==> upt(i,j) = Cons(i,upt(succ(i),j))"
 19.1092 -apply (rule trans)
 19.1093 -apply (rule upt_rec, auto)
 19.1094 -done
 19.1095 -
 19.1096 -lemma upt_type [simp,TC]: "j:nat ==> upt(i,j):list(nat)"
 19.1097 -by (induct_tac "j", auto)
 19.1098 -
 19.1099 -(*LOOPS as a simprule, since j<=j*)
 19.1100 -lemma upt_add_eq_append:
 19.1101 -     "[| i le j; j:nat; k:nat |] ==> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)"
 19.1102 -apply (induct_tac "k")
 19.1103 -apply (auto simp add: app_assoc app_type)
 19.1104 -apply (rule_tac j = j in le_trans, auto)
 19.1105 -done
 19.1106 -
 19.1107 -lemma length_upt [simp]: "[| i:nat; j:nat |] ==>length(upt(i,j)) = j #- i"
 19.1108 -apply (induct_tac "j")
 19.1109 -apply (rule_tac [2] sym)
 19.1110 -apply (auto dest!: not_lt_imp_le simp add: diff_succ diff_is_0_iff)
 19.1111 -done
 19.1112 -
 19.1113 -lemma nth_upt [rule_format,simp]:
 19.1114 -     "[| i:nat; j:nat; k:nat |] ==> i #+ k < j --> nth(k, upt(i,j)) = i #+ k"
 19.1115 -apply (induct_tac "j", simp)
 19.1116 -apply (simp add: nth_append le_iff)
 19.1117 -apply (auto dest!: not_lt_imp_le 
 19.1118 -            simp add: nth_append less_diff_conv add_commute)
 19.1119 -done
 19.1120 -
 19.1121 -lemma take_upt [rule_format,simp]:
 19.1122 -     "[| m:nat; n:nat |] ==>
 19.1123 -         \<forall>i \<in> nat. i #+ m le n --> take(m, upt(i,n)) = upt(i,i#+m)"
 19.1124 -apply (induct_tac "m")
 19.1125 -apply (simp (no_asm_simp) add: take_0)
 19.1126 -apply clarify
 19.1127 -apply (subst upt_rec, simp) 
 19.1128 -apply (rule sym)
 19.1129 -apply (subst upt_rec, simp) 
 19.1130 -apply (simp_all del: upt.simps)
 19.1131 -apply (rule_tac j = "succ (i #+ x) " in lt_trans2)
 19.1132 -apply auto
 19.1133 -done
 19.1134 -
 19.1135 -lemma map_succ_upt:
 19.1136 -     "[| m:nat; n:nat |] ==> map(succ, upt(m,n))= upt(succ(m), succ(n))"
 19.1137 -apply (induct_tac "n")
 19.1138 -apply (auto simp add: map_app_distrib)
 19.1139 -done
 19.1140 -
 19.1141 -lemma nth_map [rule_format,simp]:
 19.1142 -     "xs:list(A) ==>
 19.1143 -      \<forall>n \<in> nat. n < length(xs) --> nth(n, map(f, xs)) = f(nth(n, xs))"
 19.1144 -apply (induct_tac "xs", simp)
 19.1145 -apply (rule ballI)
 19.1146 -apply (induct_tac "n", auto)
 19.1147 -done
 19.1148 -
 19.1149 -lemma nth_map_upt [rule_format]:
 19.1150 -     "[| m:nat; n:nat |] ==>
 19.1151 -      \<forall>i \<in> nat. i < n #- m --> nth(i, map(f, upt(m,n))) = f(m #+ i)"
 19.1152 -apply (rule_tac n = m and m = n in diff_induct, typecheck, simp, simp) 
 19.1153 -apply (subst map_succ_upt [symmetric], simp_all, clarify) 
 19.1154 -apply (subgoal_tac "i < length (upt (0, x))")
 19.1155 - prefer 2 
 19.1156 - apply (simp add: less_diff_conv) 
 19.1157 - apply (rule_tac j = "succ (i #+ y) " in lt_trans2)
 19.1158 -  apply simp 
 19.1159 - apply simp 
 19.1160 -apply (subgoal_tac "i < length (upt (y, x))")
 19.1161 - apply (simp_all add: add_commute less_diff_conv)
 19.1162 -done
 19.1163 -
 19.1164 -(** sublist (a generalization of nth to sets) **)
 19.1165 -
 19.1166 -definition
 19.1167 -  sublist :: "[i, i] => i"  where
 19.1168 -    "sublist(xs, A)==
 19.1169 -     map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))"
 19.1170 -
 19.1171 -lemma sublist_0 [simp]: "xs:list(A) ==>sublist(xs, 0) =Nil"
 19.1172 -by (unfold sublist_def, auto)
 19.1173 -
 19.1174 -lemma sublist_Nil [simp]: "sublist(Nil, A) = Nil"
 19.1175 -by (unfold sublist_def, auto)
 19.1176 -
 19.1177 -lemma sublist_shift_lemma:
 19.1178 - "[| xs:list(B); i:nat |] ==>
 19.1179 -  map(fst, filter(%p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) =
 19.1180 -  map(fst, filter(%p. snd(p):nat & snd(p) #+ i:A, zip(xs,upt(0,length(xs)))))"
 19.1181 -apply (erule list_append_induct)
 19.1182 -apply (simp (no_asm_simp))
 19.1183 -apply (auto simp add: add_commute length_app filter_append map_app_distrib)
 19.1184 -done
 19.1185 -
 19.1186 -lemma sublist_type [simp,TC]:
 19.1187 -     "xs:list(B) ==> sublist(xs, A):list(B)"
 19.1188 -apply (unfold sublist_def)
 19.1189 -apply (induct_tac "xs")
 19.1190 -apply (auto simp add: filter_append map_app_distrib)
 19.1191 -done
 19.1192 -
 19.1193 -lemma upt_add_eq_append2:
 19.1194 -     "[| i:nat; j:nat |] ==> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)"
 19.1195 -by (simp add: upt_add_eq_append [of 0] nat_0_le)
 19.1196 -
 19.1197 -lemma sublist_append:
 19.1198 - "[| xs:list(B); ys:list(B)  |] ==>
 19.1199 -  sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j:nat. j #+ length(xs): A})"
 19.1200 -apply (unfold sublist_def)
 19.1201 -apply (erule_tac l = ys in list_append_induct, simp)
 19.1202 -apply (simp (no_asm_simp) add: upt_add_eq_append2 app_assoc [symmetric])
 19.1203 -apply (auto simp add: sublist_shift_lemma length_type map_app_distrib app_assoc)
 19.1204 -apply (simp_all add: add_commute)
 19.1205 -done
 19.1206 -
 19.1207 -
 19.1208 -lemma sublist_Cons:
 19.1209 -     "[| xs:list(B); x:B |] ==>
 19.1210 -      sublist(Cons(x, xs), A) = 
 19.1211 -      (if 0:A then [x] else []) @ sublist(xs, {j:nat. succ(j) : A})"
 19.1212 -apply (erule_tac l = xs in list_append_induct)
 19.1213 -apply (simp (no_asm_simp) add: sublist_def)  
 19.1214 -apply (simp del: app_Cons add: app_Cons [symmetric] sublist_append, simp) 
 19.1215 -done
 19.1216 -
 19.1217 -lemma sublist_singleton [simp]:
 19.1218 -     "sublist([x], A) = (if 0 : A then [x] else [])"
 19.1219 -by (simp add: sublist_Cons)
 19.1220 -
 19.1221 -lemma sublist_upt_eq_take [rule_format, simp]:
 19.1222 -    "xs:list(A) ==> ALL n:nat. sublist(xs,n) = take(n,xs)"
 19.1223 -apply (erule list.induct, simp) 
 19.1224 -apply (clarify ); 
 19.1225 -apply (erule natE) 
 19.1226 -apply (simp_all add: nat_eq_Collect_lt Ord_mem_iff_lt sublist_Cons)
 19.1227 -done
 19.1228 -
 19.1229 -lemma sublist_Int_eq:
 19.1230 -     "xs : list(B) ==> sublist(xs, A \<inter> nat) = sublist(xs, A)"
 19.1231 -apply (erule list.induct)
 19.1232 -apply (simp_all add: sublist_Cons) 
 19.1233 -done
 19.1234 -
 19.1235 -text{*Repetition of a List Element*}
 19.1236 -
 19.1237 -consts   repeat :: "[i,i]=>i"
 19.1238 -primrec
 19.1239 -  "repeat(a,0) = []"
 19.1240 -
 19.1241 -  "repeat(a,succ(n)) = Cons(a,repeat(a,n))"
 19.1242 -
 19.1243 -lemma length_repeat: "n \<in> nat ==> length(repeat(a,n)) = n"
 19.1244 -by (induct_tac n, auto)
 19.1245 -
 19.1246 -lemma repeat_succ_app: "n \<in> nat ==> repeat(a,succ(n)) = repeat(a,n) @ [a]"
 19.1247 -apply (induct_tac n)
 19.1248 -apply (simp_all del: app_Cons add: app_Cons [symmetric])
 19.1249 -done
 19.1250 -
 19.1251 -lemma repeat_type [TC]: "[|a \<in> A; n \<in> nat|] ==> repeat(a,n) \<in> list(A)"
 19.1252 -by (induct_tac n, auto)
 19.1253 -
 19.1254 -end
    20.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    20.2 +++ b/src/ZF/List_ZF.thy	Mon Feb 11 15:40:21 2008 +0100
    20.3 @@ -0,0 +1,1251 @@
    20.4 +(*  Title:      ZF/List
    20.5 +    ID:         $Id$
    20.6 +    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    20.7 +    Copyright   1994  University of Cambridge
    20.8 +
    20.9 +*)
   20.10 +
   20.11 +header{*Lists in Zermelo-Fraenkel Set Theory*}
   20.12 +
   20.13 +theory List_ZF imports Datatype_ZF ArithSimp begin
   20.14 +
   20.15 +consts
   20.16 +  list       :: "i=>i"
   20.17 +
   20.18 +datatype
   20.19 +  "list(A)" = Nil | Cons ("a:A", "l: list(A)")
   20.20 +
   20.21 +
   20.22 +syntax
   20.23 + "[]"        :: i                                       ("[]")
   20.24 + "@List"     :: "is => i"                                 ("[(_)]")
   20.25 +
   20.26 +translations
   20.27 +  "[x, xs]"     == "Cons(x, [xs])"
   20.28 +  "[x]"         == "Cons(x, [])"
   20.29 +  "[]"          == "Nil"
   20.30 +
   20.31 +
   20.32 +consts
   20.33 +  length :: "i=>i"
   20.34 +  hd     :: "i=>i"
   20.35 +  tl     :: "i=>i"
   20.36 +
   20.37 +primrec
   20.38 +  "length([]) = 0"
   20.39 +  "length(Cons(a,l)) = succ(length(l))"
   20.40 +
   20.41 +primrec
   20.42 +  "hd([]) = 0"
   20.43 +  "hd(Cons(a,l)) = a"
   20.44 +
   20.45 +primrec
   20.46 +  "tl([]) = []"
   20.47 +  "tl(Cons(a,l)) = l"
   20.48 +
   20.49 +
   20.50 +consts
   20.51 +  map         :: "[i=>i, i] => i"
   20.52 +  set_of_list :: "i=>i"
   20.53 +  app         :: "[i,i]=>i"                        (infixr "@" 60)
   20.54 +
   20.55 +(*map is a binding operator -- it applies to meta-level functions, not
   20.56 +object-level functions.  This simplifies the final form of term_rec_conv,
   20.57 +although complicating its derivation.*)
   20.58 +primrec
   20.59 +  "map(f,[]) = []"
   20.60 +  "map(f,Cons(a,l)) = Cons(f(a), map(f,l))"
   20.61 +
   20.62 +primrec
   20.63 +  "set_of_list([]) = 0"
   20.64 +  "set_of_list(Cons(a,l)) = cons(a, set_of_list(l))"
   20.65 +
   20.66 +primrec
   20.67 +  app_Nil:  "[] @ ys = ys"
   20.68 +  app_Cons: "(Cons(a,l)) @ ys = Cons(a, l @ ys)"
   20.69 +
   20.70 +
   20.71 +consts
   20.72 +  rev :: "i=>i"
   20.73 +  flat       :: "i=>i"
   20.74 +  list_add   :: "i=>i"
   20.75 +
   20.76 +primrec
   20.77 +  "rev([]) = []"
   20.78 +  "rev(Cons(a,l)) = rev(l) @ [a]"
   20.79 +
   20.80 +primrec
   20.81 +  "flat([]) = []"
   20.82 +  "flat(Cons(l,ls)) = l @ flat(ls)"
   20.83 +
   20.84 +primrec
   20.85 +  "list_add([]) = 0"
   20.86 +  "list_add(Cons(a,l)) = a #+ list_add(l)"
   20.87 +
   20.88 +consts
   20.89 +  drop       :: "[i,i]=>i"
   20.90 +
   20.91 +primrec
   20.92 +  drop_0:    "drop(0,l) = l"
   20.93 +  drop_succ: "drop(succ(i), l) = tl (drop(i,l))"
   20.94 +
   20.95 +
   20.96 +(*** Thanks to Sidi Ehmety for the following ***)
   20.97 +
   20.98 +definition
   20.99 +(* Function `take' returns the first n elements of a list *)
  20.100 +  take     :: "[i,i]=>i"  where
  20.101 +  "take(n, as) == list_rec(lam n:nat. [],
  20.102 +		%a l r. lam n:nat. nat_case([], %m. Cons(a, r`m), n), as)`n"
  20.103 +
  20.104 +definition
  20.105 +  nth :: "[i, i]=>i"  where
  20.106 +  --{*returns the (n+1)th element of a list, or 0 if the
  20.107 +   list is too short.*}
  20.108 +  "nth(n, as) == list_rec(lam n:nat. 0,
  20.109 + 		          %a l r. lam n:nat. nat_case(a, %m. r`m, n), as) ` n"
  20.110 +
  20.111 +definition
  20.112 +  list_update :: "[i, i, i]=>i"  where
  20.113 +  "list_update(xs, i, v) == list_rec(lam n:nat. Nil,
  20.114 +      %u us vs. lam n:nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i"
  20.115 +
  20.116 +consts
  20.117 +  filter :: "[i=>o, i] => i"
  20.118 +  upt :: "[i, i] =>i"
  20.119 +
  20.120 +primrec
  20.121 +  "filter(P, Nil) = Nil"
  20.122 +  "filter(P, Cons(x, xs)) =
  20.123 +     (if P(x) then Cons(x, filter(P, xs)) else filter(P, xs))"
  20.124 +
  20.125 +primrec
  20.126 +  "upt(i, 0) = Nil"
  20.127 +  "upt(i, succ(j)) = (if i le j then upt(i, j)@[j] else Nil)"
  20.128 +
  20.129 +definition
  20.130 +  min :: "[i,i] =>i"  where
  20.131 +    "min(x, y) == (if x le y then x else y)"
  20.132 +
  20.133 +definition
  20.134 +  max :: "[i, i] =>i"  where
  20.135 +    "max(x, y) == (if x le y then y else x)"
  20.136 +
  20.137 +(*** Aspects of the datatype definition ***)
  20.138 +
  20.139 +declare list.intros [simp,TC]
  20.140 +
  20.141 +(*An elimination rule, for type-checking*)
  20.142 +inductive_cases ConsE: "Cons(a,l) : list(A)"
  20.143 +
  20.144 +lemma Cons_type_iff [simp]: "Cons(a,l) \<in> list(A) <-> a \<in> A & l \<in> list(A)"
  20.145 +by (blast elim: ConsE) 
  20.146 +
  20.147 +(*Proving freeness results*)
  20.148 +lemma Cons_iff: "Cons(a,l)=Cons(a',l') <-> a=a' & l=l'"
  20.149 +by auto
  20.150 +
  20.151 +lemma Nil_Cons_iff: "~ Nil=Cons(a,l)"
  20.152 +by auto
  20.153 +
  20.154 +lemma list_unfold: "list(A) = {0} + (A * list(A))"
  20.155 +by (blast intro!: list.intros [unfolded list.con_defs]
  20.156 +          elim: list.cases [unfolded list.con_defs])
  20.157 +
  20.158 +
  20.159 +(**  Lemmas to justify using "list" in other recursive type definitions **)
  20.160 +
  20.161 +lemma list_mono: "A<=B ==> list(A) <= list(B)"
  20.162 +apply (unfold list.defs )
  20.163 +apply (rule lfp_mono)
  20.164 +apply (simp_all add: list.bnd_mono)
  20.165 +apply (assumption | rule univ_mono basic_monos)+
  20.166 +done
  20.167 +
  20.168 +(*There is a similar proof by list induction.*)
  20.169 +lemma list_univ: "list(univ(A)) <= univ(A)"
  20.170 +apply (unfold list.defs list.con_defs)
  20.171 +apply (rule lfp_lowerbound)
  20.172 +apply (rule_tac [2] A_subset_univ [THEN univ_mono])
  20.173 +apply (blast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ)
  20.174 +done
  20.175 +
  20.176 +(*These two theorems justify datatypes involving list(nat), list(A), ...*)
  20.177 +lemmas list_subset_univ = subset_trans [OF list_mono list_univ]
  20.178 +
  20.179 +lemma list_into_univ: "[| l: list(A);  A <= univ(B) |] ==> l: univ(B)"
  20.180 +by (blast intro: list_subset_univ [THEN subsetD])
  20.181 +
  20.182 +lemma list_case_type:
  20.183 +    "[| l: list(A);
  20.184 +        c: C(Nil);
  20.185 +        !!x y. [| x: A;  y: list(A) |] ==> h(x,y): C(Cons(x,y))
  20.186 +     |] ==> list_case(c,h,l) : C(l)"
  20.187 +by (erule list.induct, auto)
  20.188 +
  20.189 +lemma list_0_triv: "list(0) = {Nil}"
  20.190 +apply (rule equalityI, auto) 
  20.191 +apply (induct_tac x, auto) 
  20.192 +done
  20.193 +
  20.194 +
  20.195 +(*** List functions ***)
  20.196 +
  20.197 +lemma tl_type: "l: list(A) ==> tl(l) : list(A)"
  20.198 +apply (induct_tac "l")
  20.199 +apply (simp_all (no_asm_simp) add: list.intros)
  20.200 +done
  20.201 +
  20.202 +(** drop **)
  20.203 +
  20.204 +lemma drop_Nil [simp]: "i:nat ==> drop(i, Nil) = Nil"
  20.205 +apply (induct_tac "i")
  20.206 +apply (simp_all (no_asm_simp))
  20.207 +done
  20.208 +
  20.209 +lemma drop_succ_Cons [simp]: "i:nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)"
  20.210 +apply (rule sym)
  20.211 +apply (induct_tac "i")
  20.212 +apply (simp (no_asm))
  20.213 +apply (simp (no_asm_simp))
  20.214 +done
  20.215 +
  20.216 +lemma drop_type [simp,TC]: "[| i:nat; l: list(A) |] ==> drop(i,l) : list(A)"
  20.217 +apply (induct_tac "i")
  20.218 +apply (simp_all (no_asm_simp) add: tl_type)
  20.219 +done
  20.220 +
  20.221 +declare drop_succ [simp del]
  20.222 +
  20.223 +
  20.224 +(** Type checking -- proved by induction, as usual **)
  20.225 +
  20.226 +lemma list_rec_type [TC]:
  20.227 +    "[| l: list(A);
  20.228 +        c: C(Nil);
  20.229 +        !!x y r. [| x:A;  y: list(A);  r: C(y) |] ==> h(x,y,r): C(Cons(x,y))
  20.230 +     |] ==> list_rec(c,h,l) : C(l)"
  20.231 +by (induct_tac "l", auto)
  20.232 +
  20.233 +(** map **)
  20.234 +
  20.235 +lemma map_type [TC]:
  20.236 +    "[| l: list(A);  !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)"
  20.237 +apply (simp add: map_list_def)
  20.238 +apply (typecheck add: list.intros list_rec_type, blast)
  20.239 +done
  20.240 +
  20.241 +lemma map_type2 [TC]: "l: list(A) ==> map(h,l) : list({h(u). u:A})"
  20.242 +apply (erule map_type)
  20.243 +apply (erule RepFunI)
  20.244 +done
  20.245 +
  20.246 +(** length **)
  20.247 +
  20.248 +lemma length_type [TC]: "l: list(A) ==> length(l) : nat"
  20.249 +by (simp add: length_list_def)
  20.250 +
  20.251 +lemma lt_length_in_nat:
  20.252 +   "[|x < length(xs); xs \<in> list(A)|] ==> x \<in> nat"
  20.253 +by (frule lt_nat_in_nat, typecheck) 
  20.254 +
  20.255 +(** app **)
  20.256 +
  20.257 +lemma app_type [TC]: "[| xs: list(A);  ys: list(A) |] ==> xs@ys : list(A)"
  20.258 +by (simp add: app_list_def)
  20.259 +
  20.260 +(** rev **)
  20.261 +
  20.262 +lemma rev_type [TC]: "xs: list(A) ==> rev(xs) : list(A)"
  20.263 +by (simp add: rev_list_def)
  20.264 +
  20.265 +
  20.266 +(** flat **)
  20.267 +
  20.268 +lemma flat_type [TC]: "ls: list(list(A)) ==> flat(ls) : list(A)"
  20.269 +by (simp add: flat_list_def)
  20.270 +
  20.271 +
  20.272 +(** set_of_list **)
  20.273 +
  20.274 +lemma set_of_list_type [TC]: "l: list(A) ==> set_of_list(l) : Pow(A)"
  20.275 +apply (unfold set_of_list_list_def)
  20.276 +apply (erule list_rec_type, auto)
  20.277 +done
  20.278 +
  20.279 +lemma set_of_list_append:
  20.280 +     "xs: list(A) ==> set_of_list (xs@ys) = set_of_list(xs) Un set_of_list(ys)"
  20.281 +apply (erule list.induct)
  20.282 +apply (simp_all (no_asm_simp) add: Un_cons)
  20.283 +done
  20.284 +
  20.285 +
  20.286 +(** list_add **)
  20.287 +
  20.288 +lemma list_add_type [TC]: "xs: list(nat) ==> list_add(xs) : nat"
  20.289 +by (simp add: list_add_list_def)
  20.290 +
  20.291 +
  20.292 +(*** theorems about map ***)
  20.293 +
  20.294 +lemma map_ident [simp]: "l: list(A) ==> map(%u. u, l) = l"
  20.295 +apply (induct_tac "l")
  20.296 +apply (simp_all (no_asm_simp))
  20.297 +done
  20.298 +
  20.299 +lemma map_compose: "l: list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)"
  20.300 +apply (induct_tac "l")
  20.301 +apply (simp_all (no_asm_simp))
  20.302 +done
  20.303 +
  20.304 +lemma map_app_distrib: "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)"
  20.305 +apply (induct_tac "xs")
  20.306 +apply (simp_all (no_asm_simp))
  20.307 +done
  20.308 +
  20.309 +lemma map_flat: "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))"
  20.310 +apply (induct_tac "ls")
  20.311 +apply (simp_all (no_asm_simp) add: map_app_distrib)
  20.312 +done
  20.313 +
  20.314 +lemma list_rec_map:
  20.315 +     "l: list(A) ==>
  20.316 +      list_rec(c, d, map(h,l)) =
  20.317 +      list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)"
  20.318 +apply (induct_tac "l")
  20.319 +apply (simp_all (no_asm_simp))
  20.320 +done
  20.321 +
  20.322 +(** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **)
  20.323 +
  20.324 +(* c : list(Collect(B,P)) ==> c : list(B) *)
  20.325 +lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD, standard]
  20.326 +
  20.327 +lemma map_list_Collect: "l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)"
  20.328 +apply (induct_tac "l")
  20.329 +apply (simp_all (no_asm_simp))
  20.330 +done
  20.331 +
  20.332 +(*** theorems about length ***)
  20.333 +
  20.334 +lemma length_map [simp]: "xs: list(A) ==> length(map(h,xs)) = length(xs)"
  20.335 +by (induct_tac "xs", simp_all)
  20.336 +
  20.337 +lemma length_app [simp]:
  20.338 +     "[| xs: list(A); ys: list(A) |]
  20.339 +      ==> length(xs@ys) = length(xs) #+ length(ys)"
  20.340 +by (induct_tac "xs", simp_all)
  20.341 +
  20.342 +lemma length_rev [simp]: "xs: list(A) ==> length(rev(xs)) = length(xs)"
  20.343 +apply (induct_tac "xs")
  20.344 +apply (simp_all (no_asm_simp) add: length_app)
  20.345 +done
  20.346 +
  20.347 +lemma length_flat:
  20.348 +     "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))"
  20.349 +apply (induct_tac "ls")
  20.350 +apply (simp_all (no_asm_simp) add: length_app)
  20.351 +done
  20.352 +
  20.353 +(** Length and drop **)
  20.354 +
  20.355 +(*Lemma for the inductive step of drop_length*)
  20.356 +lemma drop_length_Cons [rule_format]:
  20.357 +     "xs: list(A) ==>
  20.358 +           \<forall>x.  EX z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)"
  20.359 +by (erule list.induct, simp_all)
  20.360 +
  20.361 +lemma drop_length [rule_format]:
  20.362 +     "l: list(A) ==> \<forall>i \<in> length(l). (EX z zs. drop(i,l) = Cons(z,zs))"
  20.363 +apply (erule list.induct, simp_all, safe)
  20.364 +apply (erule drop_length_Cons)
  20.365 +apply (rule natE)
  20.366 +apply (erule Ord_trans [OF asm_rl length_type Ord_nat], assumption, simp_all)
  20.367 +apply (blast intro: succ_in_naturalD length_type)
  20.368 +done
  20.369 +
  20.370 +
  20.371 +(*** theorems about app ***)
  20.372 +
  20.373 +lemma app_right_Nil [simp]: "xs: list(A) ==> xs@Nil=xs"
  20.374 +by (erule list.induct, simp_all)
  20.375 +
  20.376 +lemma app_assoc: "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)"
  20.377 +by (induct_tac "xs", simp_all)
  20.378 +
  20.379 +lemma flat_app_distrib: "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)"
  20.380 +apply (induct_tac "ls")
  20.381 +apply (simp_all (no_asm_simp) add: app_assoc)
  20.382 +done
  20.383 +
  20.384 +(*** theorems about rev ***)
  20.385 +
  20.386 +lemma rev_map_distrib: "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))"
  20.387 +apply (induct_tac "l")
  20.388 +apply (simp_all (no_asm_simp) add: map_app_distrib)
  20.389 +done
  20.390 +
  20.391 +(*Simplifier needs the premises as assumptions because rewriting will not
  20.392 +  instantiate the variable ?A in the rules' typing conditions; note that
  20.393 +  rev_type does not instantiate ?A.  Only the premises do.
  20.394 +*)
  20.395 +lemma rev_app_distrib:
  20.396 +     "[| xs: list(A);  ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"
  20.397 +apply (erule list.induct)
  20.398 +apply (simp_all add: app_assoc)
  20.399 +done
  20.400 +
  20.401 +lemma rev_rev_ident [simp]: "l: list(A) ==> rev(rev(l))=l"
  20.402 +apply (induct_tac "l")
  20.403 +apply (simp_all (no_asm_simp) add: rev_app_distrib)
  20.404 +done
  20.405 +
  20.406 +lemma rev_flat: "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))"
  20.407 +apply (induct_tac "ls")
  20.408 +apply (simp_all add: map_app_distrib flat_app_distrib rev_app_distrib)
  20.409 +done
  20.410 +
  20.411 +
  20.412 +(*** theorems about list_add ***)
  20.413 +
  20.414 +lemma list_add_app:
  20.415 +     "[| xs: list(nat);  ys: list(nat) |]
  20.416 +      ==> list_add(xs@ys) = list_add(ys) #+ list_add(xs)"
  20.417 +apply (induct_tac "xs", simp_all)
  20.418 +done
  20.419 +
  20.420 +lemma list_add_rev: "l: list(nat) ==> list_add(rev(l)) = list_add(l)"
  20.421 +apply (induct_tac "l")
  20.422 +apply (simp_all (no_asm_simp) add: list_add_app)
  20.423 +done
  20.424 +
  20.425 +lemma list_add_flat:
  20.426 +     "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))"
  20.427 +apply (induct_tac "ls")
  20.428 +apply (simp_all (no_asm_simp) add: list_add_app)
  20.429 +done
  20.430 +
  20.431 +(** New induction rules **)
  20.432 +
  20.433 +lemma list_append_induct [case_names Nil snoc, consumes 1]:
  20.434 +    "[| l: list(A);
  20.435 +        P(Nil);
  20.436 +        !!x y. [| x: A;  y: list(A);  P(y) |] ==> P(y @ [x])
  20.437 +     |] ==> P(l)"
  20.438 +apply (subgoal_tac "P(rev(rev(l)))", simp)
  20.439 +apply (erule rev_type [THEN list.induct], simp_all)
  20.440 +done
  20.441 +
  20.442 +lemma list_complete_induct_lemma [rule_format]:
  20.443 + assumes ih: 
  20.444 +    "\<And>l. [| l \<in> list(A); 
  20.445 +             \<forall>l' \<in> list(A). length(l') < length(l) --> P(l')|] 
  20.446 +          ==> P(l)"
  20.447 +  shows "n \<in> nat ==> \<forall>l \<in> list(A). length(l) < n --> P(l)"
  20.448 +apply (induct_tac n, simp)
  20.449 +apply (blast intro: ih elim!: leE) 
  20.450 +done
  20.451 +
  20.452 +theorem list_complete_induct:
  20.453 +      "[| l \<in> list(A); 
  20.454 +          \<And>l. [| l \<in> list(A); 
  20.455 +                  \<forall>l' \<in> list(A). length(l') < length(l) --> P(l')|] 
  20.456 +               ==> P(l)
  20.457 +       |] ==> P(l)"
  20.458 +apply (rule list_complete_induct_lemma [of A]) 
  20.459 +   prefer 4 apply (rule le_refl, simp) 
  20.460 +  apply blast 
  20.461 + apply simp 
  20.462 +apply assumption
  20.463 +done
  20.464 +
  20.465 +
  20.466 +(*** Thanks to Sidi Ehmety for these results about min, take, etc. ***)
  20.467 +
  20.468 +(** min FIXME: replace by Int! **)
  20.469 +(* Min theorems are also true for i, j ordinals *)
  20.470 +lemma min_sym: "[| i:nat; j:nat |] ==> min(i,j)=min(j,i)"
  20.471 +apply (unfold min_def)
  20.472 +apply (auto dest!: not_lt_imp_le dest: lt_not_sym intro: le_anti_sym)
  20.473 +done
  20.474 +
  20.475 +lemma min_type [simp,TC]: "[| i:nat; j:nat |] ==> min(i,j):nat"
  20.476 +by (unfold min_def, auto)
  20.477 +
  20.478 +lemma min_0 [simp]: "i:nat ==> min(0,i) = 0"
  20.479 +apply (unfold min_def)
  20.480 +apply (auto dest: not_lt_imp_le)
  20.481 +done
  20.482 +
  20.483 +lemma min_02 [simp]: "i:nat ==> min(i, 0) = 0"
  20.484 +apply (unfold min_def)
  20.485 +apply (auto dest: not_lt_imp_le)
  20.486 +done
  20.487 +
  20.488 +lemma lt_min_iff: "[| i:nat; j:nat; k:nat |] ==> i<min(j,k) <-> i<j & i<k"
  20.489 +apply (unfold min_def)
  20.490 +apply (auto dest!: not_lt_imp_le intro: lt_trans2 lt_trans)
  20.491 +done
  20.492 +
  20.493 +lemma min_succ_succ [simp]:
  20.494 +     "[| i:nat; j:nat |] ==>  min(succ(i), succ(j))= succ(min(i, j))"
  20.495 +apply (unfold min_def, auto)
  20.496 +done
  20.497 +
  20.498 +(*** more theorems about lists ***)
  20.499 +
  20.500 +(** filter **)
  20.501 +
  20.502 +lemma filter_append [simp]:
  20.503 +     "xs:list(A) ==> filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)"
  20.504 +by (induct_tac "xs", auto)
  20.505 +
  20.506 +lemma filter_type [simp,TC]: "xs:list(A) ==> filter(P, xs):list(A)"
  20.507 +by (induct_tac "xs", auto)
  20.508 +
  20.509 +lemma length_filter: "xs:list(A) ==> length(filter(P, xs)) le length(xs)"
  20.510 +apply (induct_tac "xs", auto)
  20.511 +apply (rule_tac j = "length (l) " in le_trans)
  20.512 +apply (auto simp add: le_iff)
  20.513 +done
  20.514 +
  20.515 +lemma filter_is_subset: "xs:list(A) ==> set_of_list(filter(P,xs)) <= set_of_list(xs)"
  20.516 +by (induct_tac "xs", auto)
  20.517 +
  20.518 +lemma filter_False [simp]: "xs:list(A) ==> filter(%p. False, xs) = Nil"
  20.519 +by (induct_tac "xs", auto)
  20.520 +
  20.521 +lemma filter_True [simp]: "xs:list(A) ==> filter(%p. True, xs) = xs"
  20.522 +by (induct_tac "xs", auto)
  20.523 +
  20.524 +(** length **)
  20.525 +
  20.526 +lemma length_is_0_iff [simp]: "xs:list(A) ==> length(xs)=0 <-> xs=Nil"
  20.527 +by (erule list.induct, auto)
  20.528 +
  20.529 +lemma length_is_0_iff2 [simp]: "xs:list(A) ==> 0 = length(xs) <-> xs=Nil"
  20.530 +by (erule list.induct, auto)
  20.531 +
  20.532 +lemma length_tl [simp]: "xs:list(A) ==> length(tl(xs)) = length(xs) #- 1"
  20.533 +by (erule list.induct, auto)
  20.534 +
  20.535 +lemma length_greater_0_iff: "xs:list(A) ==> 0<length(xs) <-> xs ~= Nil"
  20.536 +by (erule list.induct, auto)
  20.537 +
  20.538 +lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) <-> (EX y ys. xs=Cons(y, ys) & length(ys)=n)"
  20.539 +by (erule list.induct, auto)
  20.540 +
  20.541 +(** more theorems about append **)
  20.542 +
  20.543 +lemma append_is_Nil_iff [simp]:
  20.544 +     "xs:list(A) ==> (xs@ys = Nil) <-> (xs=Nil & ys = Nil)"
  20.545 +by (erule list.induct, auto)
  20.546 +
  20.547 +lemma append_is_Nil_iff2 [simp]:
  20.548 +     "xs:list(A) ==> (Nil = xs@ys) <-> (xs=Nil & ys = Nil)"
  20.549 +by (erule list.induct, auto)
  20.550 +
  20.551 +lemma append_left_is_self_iff [simp]:
  20.552 +     "xs:list(A) ==> (xs@ys = xs) <-> (ys = Nil)"
  20.553 +by (erule list.induct, auto)
  20.554 +
  20.555 +lemma append_left_is_self_iff2 [simp]:
  20.556 +     "xs:list(A) ==> (xs = xs@ys) <-> (ys = Nil)"
  20.557 +by (erule list.induct, auto)
  20.558 +
  20.559 +(*TOO SLOW as a default simprule!*)
  20.560 +lemma append_left_is_Nil_iff [rule_format]:
  20.561 +     "[| xs:list(A); ys:list(A); zs:list(A) |] ==>
  20.562 +   length(ys)=length(zs) --> (xs@ys=zs <-> (xs=Nil & ys=zs))"
  20.563 +apply (erule list.induct)
  20.564 +apply (auto simp add: length_app)
  20.565 +done
  20.566 +
  20.567 +(*TOO SLOW as a default simprule!*)
  20.568 +lemma append_left_is_Nil_iff2 [rule_format]:
  20.569 +     "[| xs:list(A); ys:list(A); zs:list(A) |] ==>
  20.570 +   length(ys)=length(zs) --> (zs=ys@xs <-> (xs=Nil & ys=zs))"
  20.571 +apply (erule list.induct)
  20.572 +apply (auto simp add: length_app)
  20.573 +done
  20.574 +
  20.575 +lemma append_eq_append_iff [rule_format,simp]:
  20.576 +     "xs:list(A) ==> \<forall>ys \<in> list(A).
  20.577 +      length(xs)=length(ys) --> (xs@us = ys@vs) <-> (xs=ys & us=vs)"
  20.578 +apply (erule list.induct)
  20.579 +apply (simp (no_asm_simp))
  20.580 +apply clarify
  20.581 +apply (erule_tac a = ys in list.cases, auto)
  20.582 +done
  20.583 +
  20.584 +lemma append_eq_append [rule_format]:
  20.585 +  "xs:list(A) ==>
  20.586 +   \<forall>ys \<in> list(A). \<forall>us \<in> list(A). \<forall>vs \<in> list(A).
  20.587 +   length(us) = length(vs) --> (xs@us = ys@vs) --> (xs=ys & us=vs)"
  20.588 +apply (induct_tac "xs")
  20.589 +apply (force simp add: length_app, clarify)
  20.590 +apply (erule_tac a = ys in list.cases, simp)
  20.591 +apply (subgoal_tac "Cons (a, l) @ us =vs")
  20.592 + apply (drule rev_iffD1 [OF _ append_left_is_Nil_iff], simp_all, blast)
  20.593 +done
  20.594 +
  20.595 +lemma append_eq_append_iff2 [simp]:
  20.596 + "[| xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs) |]
  20.597 +  ==>  xs@us = ys@vs <-> (xs=ys & us=vs)"
  20.598 +apply (rule iffI)
  20.599 +apply (rule append_eq_append, auto)
  20.600 +done
  20.601 +
  20.602 +lemma append_self_iff [simp]:
  20.603 +     "[| xs:list(A); ys:list(A); zs:list(A) |] ==> xs@ys=xs@zs <-> ys=zs"
  20.604 +by simp
  20.605 +
  20.606 +lemma append_self_iff2 [simp]:
  20.607 +     "[| xs:list(A); ys:list(A); zs:list(A) |] ==> ys@xs=zs@xs <-> ys=zs"
  20.608 +by simp
  20.609 +
  20.610 +(* Can also be proved from append_eq_append_iff2,
  20.611 +but the proof requires two more hypotheses: x:A and y:A *)
  20.612 +lemma append1_eq_iff [rule_format,simp]:
  20.613 +     "xs:list(A) ==> \<forall>ys \<in> list(A). xs@[x] = ys@[y] <-> (xs = ys & x=y)"
  20.614 +apply (erule list.induct)  
  20.615 + apply clarify 
  20.616 + apply (erule list.cases)
  20.617 + apply simp_all
  20.618 +txt{*Inductive step*}  
  20.619 +apply clarify 
  20.620 +apply (erule_tac a=ys in list.cases, simp_all)  
  20.621 +done
  20.622 +
  20.623 +
  20.624 +lemma append_right_is_self_iff [simp]:
  20.625 +     "[| xs:list(A); ys:list(A) |] ==> (xs@ys = ys) <-> (xs=Nil)"
  20.626 +by (simp (no_asm_simp) add: append_left_is_Nil_iff)
  20.627 +
  20.628 +lemma append_right_is_self_iff2 [simp]:
  20.629 +     "[| xs:list(A); ys:list(A) |] ==> (ys = xs@ys) <-> (xs=Nil)"
  20.630 +apply (rule iffI)
  20.631 +apply (drule sym, auto) 
  20.632 +done
  20.633 +
  20.634 +lemma hd_append [rule_format,simp]:
  20.635 +     "xs:list(A) ==> xs ~= Nil --> hd(xs @ ys) = hd(xs)"
  20.636 +by (induct_tac "xs", auto)
  20.637 +
  20.638 +lemma tl_append [rule_format,simp]:
  20.639 +     "xs:list(A) ==> xs~=Nil --> tl(xs @ ys) = tl(xs)@ys"
  20.640 +by (induct_tac "xs", auto)
  20.641 +
  20.642 +(** rev **)
  20.643 +lemma rev_is_Nil_iff [simp]: "xs:list(A) ==> (rev(xs) = Nil <-> xs = Nil)"
  20.644 +by (erule list.induct, auto)
  20.645 +
  20.646 +lemma Nil_is_rev_iff [simp]: "xs:list(A) ==> (Nil = rev(xs) <-> xs = Nil)"
  20.647 +by (erule list.induct, auto)
  20.648 +
  20.649 +lemma rev_is_rev_iff [rule_format,simp]:
  20.650 +     "xs:list(A) ==> \<forall>ys \<in> list(A). rev(xs)=rev(ys) <-> xs=ys"
  20.651 +apply (erule list.induct, force, clarify)
  20.652 +apply (erule_tac a = ys in list.cases, auto)
  20.653 +done
  20.654 +
  20.655 +lemma rev_list_elim [rule_format]:
  20.656 +     "xs:list(A) ==>
  20.657 +      (xs=Nil --> P) --> (\<forall>ys \<in> list(A). \<forall>y \<in> A. xs =ys@[y] -->P)-->P"
  20.658 +by (erule list_append_induct, auto)
  20.659 +
  20.660 +
  20.661 +(** more theorems about drop **)
  20.662 +
  20.663 +lemma length_drop [rule_format,simp]:
  20.664 +     "n:nat ==> \<forall>xs \<in> list(A). length(drop(n, xs)) = length(xs) #- n"
  20.665 +apply (erule nat_induct)
  20.666 +apply (auto elim: list.cases)
  20.667 +done
  20.668 +
  20.669 +lemma drop_all [rule_format,simp]:
  20.670 +     "n:nat ==> \<forall>xs \<in> list(A). length(xs) le n --> drop(n, xs)=Nil"
  20.671 +apply (erule nat_induct)
  20.672 +apply (auto elim: list.cases)
  20.673 +done
  20.674 +
  20.675 +lemma drop_append [rule_format]:
  20.676 +     "n:nat ==> 
  20.677 +      \<forall>xs \<in> list(A). drop(n, xs@ys) = drop(n,xs) @ drop(n #- length(xs), ys)"
  20.678 +apply (induct_tac "n")
  20.679 +apply (auto elim: list.cases)
  20.680 +done
  20.681 +
  20.682 +lemma drop_drop:
  20.683 +    "m:nat ==> \<forall>xs \<in> list(A). \<forall>n \<in> nat. drop(n, drop(m, xs))=drop(n #+ m, xs)"
  20.684 +apply (induct_tac "m")
  20.685 +apply (auto elim: list.cases)
  20.686 +done
  20.687 +
  20.688 +(** take **)
  20.689 +
  20.690 +lemma take_0 [simp]: "xs:list(A) ==> take(0, xs) =  Nil"
  20.691 +apply (unfold take_def)
  20.692 +apply (erule list.induct, auto)
  20.693 +done
  20.694 +
  20.695 +lemma take_succ_Cons [simp]:
  20.696 +    "n:nat ==> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))"
  20.697 +by (simp add: take_def)
  20.698 +
  20.699 +(* Needed for proving take_all *)
  20.700 +lemma take_Nil [simp]: "n:nat ==> take(n, Nil) = Nil"
  20.701 +by (unfold take_def, auto)
  20.702 +
  20.703 +lemma take_all [rule_format,simp]:
  20.704 +     "n:nat ==> \<forall>xs \<in> list(A). length(xs) le n  --> take(n, xs) = xs"
  20.705 +apply (erule nat_induct)
  20.706 +apply (auto elim: list.cases) 
  20.707 +done
  20.708 +
  20.709 +lemma take_type [rule_format,simp,TC]:
  20.710 +     "xs:list(A) ==> \<forall>n \<in> nat. take(n, xs):list(A)"
  20.711 +apply (erule list.induct, simp, clarify) 
  20.712 +apply (erule natE, auto)
  20.713 +done
  20.714 +
  20.715 +lemma take_append [rule_format,simp]:
  20.716 + "xs:list(A) ==>
  20.717 +  \<forall>ys \<in> list(A). \<forall>n \<in> nat. take(n, xs @ ys) =
  20.718 +                            take(n, xs) @ take(n #- length(xs), ys)"
  20.719 +apply (erule list.induct, simp, clarify) 
  20.720 +apply (erule natE, auto)
  20.721 +done
  20.722 +
  20.723 +lemma take_take [rule_format]:
  20.724 +   "m : nat ==>
  20.725 +    \<forall>xs \<in> list(A). \<forall>n \<in> nat. take(n, take(m,xs))= take(min(n, m), xs)"
  20.726 +apply (induct_tac "m", auto)
  20.727 +apply (erule_tac a = xs in list.cases)
  20.728 +apply (auto simp add: take_Nil)
  20.729 +apply (erule_tac n=n in natE)
  20.730 +apply (auto intro: take_0 take_type)
  20.731 +done
  20.732 +
  20.733 +(** nth **)
  20.734 +
  20.735 +lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a"
  20.736 +by (simp add: nth_def) 
  20.737 +
  20.738 +lemma nth_Cons [simp]: "n:nat ==> nth(succ(n), Cons(a,l)) = nth(n,l)"
  20.739 +by (simp add: nth_def) 
  20.740 +
  20.741 +lemma nth_empty [simp]: "nth(n, Nil) = 0"
  20.742 +by (simp add: nth_def) 
  20.743 +
  20.744 +lemma nth_type [rule_format,simp,TC]:
  20.745 +     "xs:list(A) ==> \<forall>n. n < length(xs) --> nth(n,xs) : A"
  20.746 +apply (erule list.induct, simp, clarify)
  20.747 +apply (subgoal_tac "n \<in> nat")  
  20.748 + apply (erule natE, auto dest!: le_in_nat)
  20.749 +done
  20.750 +
  20.751 +lemma nth_eq_0 [rule_format]:
  20.752 +     "xs:list(A) ==> \<forall>n \<in> nat. length(xs) le n --> nth(n,xs) = 0"
  20.753 +apply (erule list.induct, simp, clarify) 
  20.754 +apply (erule natE, auto)
  20.755 +done
  20.756 +
  20.757 +lemma nth_append [rule_format]:
  20.758 +  "xs:list(A) ==> 
  20.759 +   \<forall>n \<in> nat. nth(n, xs @ ys) = (if n < length(xs) then nth(n,xs)
  20.760 +                                else nth(n #- length(xs), ys))"
  20.761 +apply (induct_tac "xs", simp, clarify) 
  20.762 +apply (erule natE, auto)
  20.763 +done
  20.764 +
  20.765 +lemma set_of_list_conv_nth:
  20.766 +    "xs:list(A)
  20.767 +     ==> set_of_list(xs) = {x:A. EX i:nat. i<length(xs) & x = nth(i,xs)}"
  20.768 +apply (induct_tac "xs", simp_all)
  20.769 +apply (rule equalityI, auto)
  20.770 +apply (rule_tac x = 0 in bexI, auto)
  20.771 +apply (erule natE, auto)
  20.772 +done
  20.773 +
  20.774 +(* Other theorems about lists *)
  20.775 +
  20.776 +lemma nth_take_lemma [rule_format]:
  20.777 + "k:nat ==>
  20.778 +  \<forall>xs \<in> list(A). (\<forall>ys \<in> list(A). k le length(xs) --> k le length(ys) -->
  20.779 +      (\<forall>i \<in> nat. i<k --> nth(i,xs) = nth(i,ys))--> take(k,xs) = take(k,ys))"
  20.780 +apply (induct_tac "k")
  20.781 +apply (simp_all (no_asm_simp) add: lt_succ_eq_0_disj all_conj_distrib)
  20.782 +apply clarify
  20.783 +(*Both lists are non-empty*)
  20.784 +apply (erule_tac a=xs in list.cases, simp) 
  20.785 +apply (erule_tac a=ys in list.cases, clarify) 
  20.786 +apply (simp (no_asm_use) )
  20.787 +apply clarify
  20.788 +apply (simp (no_asm_simp))
  20.789 +apply (rule conjI, force)
  20.790 +apply (rename_tac y ys z zs) 
  20.791 +apply (drule_tac x = zs and x1 = ys in bspec [THEN bspec], auto)   
  20.792 +done
  20.793 +
  20.794 +lemma nth_equalityI [rule_format]:
  20.795 +     "[| xs:list(A); ys:list(A); length(xs) = length(ys);
  20.796 +         \<forall>i \<in> nat. i < length(xs) --> nth(i,xs) = nth(i,ys) |]
  20.797 +      ==> xs = ys"
  20.798 +apply (subgoal_tac "length (xs) le length (ys) ")
  20.799 +apply (cut_tac k="length(xs)" and xs=xs and ys=ys in nth_take_lemma) 
  20.800 +apply (simp_all add: take_all)
  20.801 +done
  20.802 +
  20.803 +(*The famous take-lemma*)
  20.804 +
  20.805 +lemma take_equalityI [rule_format]:
  20.806 +    "[| xs:list(A); ys:list(A); (\<forall>i \<in> nat. take(i, xs) = take(i,ys)) |] 
  20.807 +     ==> xs = ys"
  20.808 +apply (case_tac "length (xs) le length (ys) ")
  20.809 +apply (drule_tac x = "length (ys) " in bspec)
  20.810 +apply (drule_tac [3] not_lt_imp_le)
  20.811 +apply (subgoal_tac [5] "length (ys) le length (xs) ")
  20.812 +apply (rule_tac [6] j = "succ (length (ys))" in le_trans)
  20.813 +apply (rule_tac [6] leI)
  20.814 +apply (drule_tac [5] x = "length (xs) " in bspec)
  20.815 +apply (simp_all add: take_all)
  20.816 +done
  20.817 +
  20.818 +lemma nth_drop [rule_format]:
  20.819 +  "n:nat ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)"
  20.820 +apply (induct_tac "n", simp_all, clarify)
  20.821 +apply (erule list.cases, auto)  
  20.822 +done
  20.823 +
  20.824 +lemma take_succ [rule_format]:
  20.825 +  "xs\<in>list(A) 
  20.826 +   ==> \<forall>i. i < length(xs) --> take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]"
  20.827 +apply (induct_tac "xs", auto)
  20.828 +apply (subgoal_tac "i\<in>nat") 
  20.829 +apply (erule natE)
  20.830 +apply (auto simp add: le_in_nat) 
  20.831 +done
  20.832 +
  20.833 +lemma take_add [rule_format]:
  20.834 +     "[|xs\<in>list(A); j\<in>nat|] 
  20.835 +      ==> \<forall>i\<in>nat. take(i #+ j, xs) = take(i,xs) @ take(j, drop(i,xs))"
  20.836 +apply (induct_tac "xs", simp_all, clarify)
  20.837 +apply (erule_tac n = i in natE, simp_all)
  20.838 +done
  20.839 +
  20.840 +lemma length_take:
  20.841 +     "l\<in>list(A) ==> \<forall>n\<in>nat. length(take(n,l)) = min(n, length(l))"
  20.842 +apply (induct_tac "l", safe, simp_all)
  20.843 +apply (erule natE, simp_all)
  20.844 +done
  20.845 +
  20.846 +subsection{*The function zip*}
  20.847 +
  20.848 +text{*Crafty definition to eliminate a type argument*}
  20.849 +
  20.850 +consts
  20.851 +  zip_aux        :: "[i,i]=>i"
  20.852 +
  20.853 +primrec (*explicit lambda is required because both arguments of "un" vary*)
  20.854 +  "zip_aux(B,[]) =
  20.855 +     (\<lambda>ys \<in> list(B). list_case([], %y l. [], ys))"
  20.856 +
  20.857 +  "zip_aux(B,Cons(x,l)) =
  20.858 +     (\<lambda>ys \<in> list(B).
  20.859 +        list_case(Nil, %y zs. Cons(<x,y>, zip_aux(B,l)`zs), ys))"
  20.860 +
  20.861 +definition
  20.862 +  zip :: "[i, i]=>i"  where
  20.863 +   "zip(xs, ys) == zip_aux(set_of_list(ys),xs)`ys"
  20.864 +
  20.865 +
  20.866 +(* zip equations *)
  20.867 +
  20.868 +lemma list_on_set_of_list: "xs \<in> list(A) ==> xs \<in> list(set_of_list(xs))"
  20.869 +apply (induct_tac xs, simp_all) 
  20.870 +apply (blast intro: list_mono [THEN subsetD]) 
  20.871 +done
  20.872 +
  20.873 +lemma zip_Nil [simp]: "ys:list(A) ==> zip(Nil, ys)=Nil"
  20.874 +apply (simp add: zip_def list_on_set_of_list [of _ A])
  20.875 +apply (erule list.cases, simp_all) 
  20.876 +done
  20.877 +
  20.878 +lemma zip_Nil2 [simp]: "xs:list(A) ==> zip(xs, Nil)=Nil"
  20.879 +apply (simp add: zip_def list_on_set_of_list [of _ A])
  20.880 +apply (erule list.cases, simp_all) 
  20.881 +done
  20.882 +
  20.883 +lemma zip_aux_unique [rule_format]:
  20.884 +     "[|B<=C;  xs \<in> list(A)|] 
  20.885 +      ==> \<forall>ys \<in> list(B). zip_aux(C,xs) ` ys = zip_aux(B,xs) ` ys"
  20.886 +apply (induct_tac xs) 
  20.887 + apply simp_all 
  20.888 + apply (blast intro: list_mono [THEN subsetD], clarify) 
  20.889 +apply (erule_tac a=ys in list.cases, auto) 
  20.890 +apply (blast intro: list_mono [THEN subsetD]) 
  20.891 +done
  20.892 +
  20.893 +lemma zip_Cons_Cons [simp]:
  20.894 +     "[| xs:list(A); ys:list(B); x:A; y:B |] ==>
  20.895 +      zip(Cons(x,xs), Cons(y, ys)) = Cons(<x,y>, zip(xs, ys))"
  20.896 +apply (simp add: zip_def, auto) 
  20.897 +apply (rule zip_aux_unique, auto) 
  20.898 +apply (simp add: list_on_set_of_list [of _ B])
  20.899 +apply (blast intro: list_on_set_of_list list_mono [THEN subsetD]) 
  20.900 +done
  20.901 +
  20.902 +lemma zip_type [rule_format,simp,TC]:
  20.903 +     "xs:list(A) ==> \<forall>ys \<in> list(B). zip(xs, ys):list(A*B)"
  20.904 +apply (induct_tac "xs")
  20.905 +apply (simp (no_asm))
  20.906 +apply clarify
  20.907 +apply (erule_tac a = ys in list.cases, auto)
  20.908 +done
  20.909 +
  20.910 +(* zip length *)
  20.911 +lemma length_zip [rule_format,simp]:
  20.912 +     "xs:list(A) ==> \<forall>ys \<in> list(B). length(zip(xs,ys)) =
  20.913 +                                     min(length(xs), length(ys))"
  20.914 +apply (unfold min_def)
  20.915 +apply (induct_tac "xs", simp_all, clarify) 
  20.916 +apply (erule_tac a = ys in list.cases, auto)
  20.917 +done
  20.918 +
  20.919 +lemma zip_append1 [rule_format]:
  20.920 + "[| ys:list(A); zs:list(B) |] ==>
  20.921 +  \<forall>xs \<in> list(A). zip(xs @ ys, zs) = 
  20.922 +                 zip(xs, take(length(xs), zs)) @ zip(ys, drop(length(xs),zs))"
  20.923 +apply (induct_tac "zs", force, clarify) 
  20.924 +apply (erule_tac a = xs in list.cases, simp_all) 
  20.925 +done
  20.926 +
  20.927 +lemma zip_append2 [rule_format]:
  20.928 + "[| xs:list(A); zs:list(B) |] ==> \<forall>ys \<in> list(B). zip(xs, ys@zs) =
  20.929 +       zip(take(length(ys), xs), ys) @ zip(drop(length(ys), xs), zs)"
  20.930 +apply (induct_tac "xs", force, clarify) 
  20.931 +apply (erule_tac a = ys in list.cases, auto)
  20.932 +done
  20.933 +
  20.934 +lemma zip_append [simp]:
  20.935 + "[| length(xs) = length(us); length(ys) = length(vs);
  20.936 +     xs:list(A); us:list(B); ys:list(A); vs:list(B) |] 
  20.937 +  ==> zip(xs@ys,us@vs) = zip(xs, us) @ zip(ys, vs)"
  20.938 +by (simp (no_asm_simp) add: zip_append1 drop_append diff_self_eq_0)
  20.939 +
  20.940 +
  20.941 +lemma zip_rev [rule_format,simp]:
  20.942 + "ys:list(B) ==> \<forall>xs \<in> list(A).
  20.943 +    length(xs) = length(ys) --> zip(rev(xs), rev(ys)) = rev(zip(xs, ys))"
  20.944 +apply (induct_tac "ys", force, clarify) 
  20.945 +apply (erule_tac a = xs in list.cases)
  20.946 +apply (auto simp add: length_rev)
  20.947 +done
  20.948 +
  20.949 +lemma nth_zip [rule_format,simp]:
  20.950 +   "ys:list(B) ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A).
  20.951 +                    i < length(xs) --> i < length(ys) -->
  20.952 +                    nth(i,zip(xs, ys)) = <nth(i,xs),nth(i, ys)>"
  20.953 +apply (induct_tac "ys", force, clarify) 
  20.954 +apply (erule_tac a = xs in list.cases, simp) 
  20.955 +apply (auto elim: natE)
  20.956 +done
  20.957 +
  20.958 +lemma set_of_list_zip [rule_format]:
  20.959 +     "[| xs:list(A); ys:list(B); i:nat |]
  20.960 +      ==> set_of_list(zip(xs, ys)) =
  20.961 +          {<x, y>:A*B. EX i:nat. i < min(length(xs), length(ys))
  20.962 +          & x = nth(i, xs) & y = nth(i, ys)}"
  20.963 +by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth)
  20.964 +
  20.965 +(** list_update **)
  20.966 +
  20.967 +lemma list_update_Nil [simp]: "i:nat ==>list_update(Nil, i, v) = Nil"
  20.968 +by (unfold list_update_def, auto)
  20.969 +
  20.970 +lemma list_update_Cons_0 [simp]: "list_update(Cons(x, xs), 0, v)= Cons(v, xs)"
  20.971 +by (unfold list_update_def, auto)
  20.972 +
  20.973 +lemma list_update_Cons_succ [simp]:
  20.974 +  "n:nat ==>
  20.975 +    list_update(Cons(x, xs), succ(n), v)= Cons(x, list_update(xs, n, v))"
  20.976 +apply (unfold list_update_def, auto)
  20.977 +done
  20.978 +
  20.979 +lemma list_update_type [rule_format,simp,TC]:
  20.980 +     "[| xs:list(A); v:A |] ==> \<forall>n \<in> nat. list_update(xs, n, v):list(A)"
  20.981 +apply (induct_tac "xs")
  20.982 +apply (simp (no_asm))
  20.983 +apply clarify
  20.984 +apply (erule natE, auto)
  20.985 +done
  20.986 +
  20.987 +lemma length_list_update [rule_format,simp]:
  20.988 +     "xs:list(A) ==> \<forall>i \<in> nat. length(list_update(xs, i, v))=length(xs)"
  20.989 +apply (induct_tac "xs")
  20.990 +apply (simp (no_asm))
  20.991 +apply clarify
  20.992 +apply (erule natE, auto)
  20.993 +done
  20.994 +
  20.995 +lemma nth_list_update [rule_format]:
  20.996 +     "[| xs:list(A) |] ==> \<forall>i \<in> nat. \<forall>j \<in> nat. i < length(xs)  -->
  20.997 +         nth(j, list_update(xs, i, x)) = (if i=j then x else nth(j, xs))"
  20.998 +apply (induct_tac "xs")
  20.999 + apply simp_all
 20.1000 +apply clarify
 20.1001 +apply (rename_tac i j) 
 20.1002 +apply (erule_tac n=i in natE) 
 20.1003 +apply (erule_tac [2] n=j in natE)
 20.1004 +apply (erule_tac n=j in natE, simp_all, force) 
 20.1005 +done
 20.1006 +
 20.1007 +lemma nth_list_update_eq [simp]:
 20.1008 +     "[| i < length(xs); xs:list(A) |] ==> nth(i, list_update(xs, i,x)) = x"
 20.1009 +by (simp (no_asm_simp) add: lt_length_in_nat nth_list_update)
 20.1010 +
 20.1011 +
 20.1012 +lemma nth_list_update_neq [rule_format,simp]:
 20.1013 +  "xs:list(A) ==> 
 20.1014 +     \<forall>i \<in> nat. \<forall>j \<in> nat. i ~= j --> nth(j, list_update(xs,i,x)) = nth(j,xs)"
 20.1015 +apply (induct_tac "xs")
 20.1016 + apply (simp (no_asm))
 20.1017 +apply clarify
 20.1018 +apply (erule natE)
 20.1019 +apply (erule_tac [2] natE, simp_all) 
 20.1020 +apply (erule natE, simp_all)
 20.1021 +done
 20.1022 +
 20.1023 +lemma list_update_overwrite [rule_format,simp]:
 20.1024 +     "xs:list(A) ==> \<forall>i \<in> nat. i < length(xs)
 20.1025 +   --> list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)"
 20.1026 +apply (induct_tac "xs")
 20.1027 + apply (simp (no_asm))
 20.1028 +apply clarify
 20.1029 +apply (erule natE, auto)
 20.1030 +done
 20.1031 +
 20.1032 +lemma list_update_same_conv [rule_format]:
 20.1033 +     "xs:list(A) ==> 
 20.1034 +      \<forall>i \<in> nat. i < length(xs) --> 
 20.1035 +                 (list_update(xs, i, x) = xs) <-> (nth(i, xs) = x)"
 20.1036 +apply (induct_tac "xs")
 20.1037 + apply (simp (no_asm))
 20.1038 +apply clarify
 20.1039 +apply (erule natE, auto)
 20.1040 +done
 20.1041 +
 20.1042 +lemma update_zip [rule_format]:
 20.1043 +     "ys:list(B) ==> 
 20.1044 +      \<forall>i \<in> nat. \<forall>xy \<in> A*B. \<forall>xs \<in> list(A).
 20.1045 +        length(xs) = length(ys) -->
 20.1046 +        list_update(zip(xs, ys), i, xy) = zip(list_update(xs, i, fst(xy)),
 20.1047 +                                              list_update(ys, i, snd(xy)))"
 20.1048 +apply (induct_tac "ys")
 20.1049 + apply auto
 20.1050 +apply (erule_tac a = xs in list.cases)
 20.1051 +apply (auto elim: natE)
 20.1052 +done
 20.1053 +
 20.1054 +lemma set_update_subset_cons [rule_format]:
 20.1055 +  "xs:list(A) ==> 
 20.1056 +   \<forall>i \<in> nat. set_of_list(list_update(xs, i, x)) <= cons(x, set_of_list(xs))"
 20.1057 +apply (induct_tac "xs")
 20.1058 + apply simp
 20.1059 +apply (rule ballI)
 20.1060 +apply (erule natE, simp_all, auto)
 20.1061 +done
 20.1062 +
 20.1063 +lemma set_of_list_update_subsetI:
 20.1064 +     "[| set_of_list(xs) <= A; xs:list(A); x:A; i:nat|]
 20.1065 +   ==> set_of_list(list_update(xs, i,x)) <= A"
 20.1066 +apply (rule subset_trans)
 20.1067 +apply (rule set_update_subset_cons, auto)
 20.1068 +done
 20.1069 +
 20.1070 +(** upt **)
 20.1071 +
 20.1072 +lemma upt_rec:
 20.1073 +     "j:nat ==> upt(i,j) = (if i<j then Cons(i, upt(succ(i), j)) else Nil)"
 20.1074 +apply (induct_tac "j", auto)
 20.1075 +apply (drule not_lt_imp_le)
 20.1076 +apply (auto simp: lt_Ord intro: le_anti_sym)
 20.1077 +done
 20.1078 +
 20.1079 +lemma upt_conv_Nil [simp]: "[| j le i; j:nat |] ==> upt(i,j) = Nil"
 20.1080 +apply (subst upt_rec, auto)
 20.1081 +apply (auto simp add: le_iff)
 20.1082 +apply (drule lt_asym [THEN notE], auto)
 20.1083 +done
 20.1084 +
 20.1085 +(*Only needed if upt_Suc is deleted from the simpset*)
 20.1086 +lemma upt_succ_append:
 20.1087 +     "[| i le j; j:nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]"
 20.1088 +by simp
 20.1089 +
 20.1090 +lemma upt_conv_Cons:
 20.1091 +     "[| i<j; j:nat |]  ==> upt(i,j) = Cons(i,upt(succ(i),j))"
 20.1092 +apply (rule trans)
 20.1093 +apply (rule upt_rec, auto)
 20.1094 +done
 20.1095 +
 20.1096 +lemma upt_type [simp,TC]: "j:nat ==> upt(i,j):list(nat)"
 20.1097 +by (induct_tac "j", auto)
 20.1098 +
 20.1099 +(*LOOPS as a simprule, since j<=j*)
 20.1100 +lemma upt_add_eq_append:
 20.1101 +     "[| i le j; j:nat; k:nat |] ==> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)"
 20.1102 +apply (induct_tac "k")
 20.1103 +apply (auto simp add: app_assoc app_type)
 20.1104 +apply (rule_tac j = j in le_trans, auto)
 20.1105 +done
 20.1106 +
 20.1107 +lemma length_upt [simp]: "[| i:nat; j:nat |] ==>length(upt(i,j)) = j #- i"
 20.1108 +apply (induct_tac "j")
 20.1109 +apply (rule_tac [2] sym)
 20.1110 +apply (auto dest!: not_lt_imp_le simp add: diff_succ diff_is_0_iff)
 20.1111 +done
 20.1112 +
 20.1113 +lemma nth_upt [rule_format,simp]:
 20.1114 +     "[| i:nat; j:nat; k:nat |] ==> i #+ k < j --> nth(k, upt(i,j)) = i #+ k"
 20.1115 +apply (induct_tac "j", simp)
 20.1116 +apply (simp add: nth_append le_iff)
 20.1117 +apply (auto dest!: not_lt_imp_le 
 20.1118 +            simp add: nth_append less_diff_conv add_commute)
 20.1119 +done
 20.1120 +
 20.1121 +lemma take_upt [rule_format,simp]:
 20.1122 +     "[| m:nat; n:nat |] ==>
 20.1123 +         \<forall>i \<in> nat. i #+ m le n --> take(m, upt(i,n)) = upt(i,i#+m)"
 20.1124 +apply (induct_tac "m")
 20.1125 +apply (simp (no_asm_simp) add: take_0)
 20.1126 +apply clarify
 20.1127 +apply (subst upt_rec, simp) 
 20.1128 +apply (rule sym)
 20.1129 +apply (subst upt_rec, simp) 
 20.1130 +apply (simp_all del: upt.simps)
 20.1131 +apply (rule_tac j = "succ (i #+ x) " in lt_trans2)
 20.1132 +apply auto
 20.1133 +done
 20.1134 +
 20.1135 +lemma map_succ_upt:
 20.1136 +     "[| m:nat; n:nat |] ==> map(succ, upt(m,n))= upt(succ(m), succ(n))"
 20.1137 +apply (induct_tac "n")
 20.1138 +apply (auto simp add: map_app_distrib)
 20.1139 +done
 20.1140 +
 20.1141 +lemma nth_map [rule_format,simp]:
 20.1142 +     "xs:list(A) ==>
 20.1143 +      \<forall>n \<in> nat. n < length(xs) --> nth(n, map(f, xs)) = f(nth(n, xs))"
 20.1144 +apply (induct_tac "xs", simp)
 20.1145 +apply (rule ballI)
 20.1146 +apply (induct_tac "n", auto)
 20.1147 +done
 20.1148 +
 20.1149 +lemma nth_map_upt [rule_format]:
 20.1150 +     "[| m:nat; n:nat |] ==>
 20.1151 +      \<forall>i \<in> nat. i < n #- m --> nth(i, map(f, upt(m,n))) = f(m #+ i)"
 20.1152 +apply (rule_tac n = m and m = n in diff_induct, typecheck, simp, simp) 
 20.1153 +apply (subst map_succ_upt [symmetric], simp_all, clarify) 
 20.1154 +apply (subgoal_tac "i < length (upt (0, x))")
 20.1155 + prefer 2 
 20.1156 + apply (simp add: less_diff_conv) 
 20.1157 + apply (rule_tac j = "succ (i #+ y) " in lt_trans2)
 20.1158 +  apply simp 
 20.1159 + apply simp 
 20.1160 +apply (subgoal_tac "i < length (upt (y, x))")
 20.1161 + apply (simp_all add: add_commute less_diff_conv)
 20.1162 +done
 20.1163 +
 20.1164 +(** sublist (a generalization of nth to sets) **)
 20.1165 +
 20.1166 +definition
 20.1167 +  sublist :: "[i, i] => i"  where
 20.1168 +    "sublist(xs, A)==
 20.1169 +     map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))"
 20.1170 +
 20.1171 +lemma sublist_0 [simp]: "xs:list(A) ==>sublist(xs, 0) =Nil"
 20.1172 +by (unfold sublist_def, auto)
 20.1173 +
 20.1174 +lemma sublist_Nil [simp]: "sublist(Nil, A) = Nil"
 20.1175 +by (unfold sublist_def, auto)
 20.1176 +
 20.1177 +lemma sublist_shift_lemma:
 20.1178 + "[| xs:list(B); i:nat |] ==>
 20.1179 +  map(fst, filter(%p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) =
 20.1180 +  map(fst, filter(%p. snd(p):nat & snd(p) #+ i:A, zip(xs,upt(0,length(xs)))))"
 20.1181 +apply (erule list_append_induct)
 20.1182 +apply (simp (no_asm_simp))
 20.1183 +apply (auto simp add: add_commute length_app filter_append map_app_distrib)
 20.1184 +done
 20.1185 +
 20.1186 +lemma sublist_type [simp,TC]:
 20.1187 +     "xs:list(B) ==> sublist(xs, A):list(B)"
 20.1188 +apply (unfold sublist_def)
 20.1189 +apply (induct_tac "xs")
 20.1190 +apply (auto simp add: filter_append map_app_distrib)
 20.1191 +done
 20.1192 +
 20.1193 +lemma upt_add_eq_append2:
 20.1194 +     "[| i:nat; j:nat |] ==> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)"
 20.1195 +by (simp add: upt_add_eq_append [of 0] nat_0_le)
 20.1196 +
 20.1197 +lemma sublist_append:
 20.1198 + "[| xs:list(B); ys:list(B)  |] ==>
 20.1199 +  sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j:nat. j #+ length(xs): A})"
 20.1200 +apply (unfold sublist_def)
 20.1201 +apply (erule_tac l = ys in list_append_induct, simp)
 20.1202 +apply (simp (no_asm_simp) add: upt_add_eq_append2 app_assoc [symmetric])
 20.1203 +apply (auto simp add: sublist_shift_lemma length_type map_app_distrib app_assoc)
 20.1204 +apply (simp_all add: add_commute)
 20.1205 +done
 20.1206 +
 20.1207 +
 20.1208 +lemma sublist_Cons:
 20.1209 +     "[| xs:list(B); x:B |] ==>
 20.1210 +      sublist(Cons(x, xs), A) = 
 20.1211 +      (if 0:A then [x] else []) @ sublist(xs, {j:nat. succ(j) : A})"
 20.1212 +apply (erule_tac l = xs in list_append_induct)
 20.1213 +apply (simp (no_asm_simp) add: sublist_def)  
 20.1214 +apply (simp del: app_Cons add: app_Cons [symmetric] sublist_append, simp) 
 20.1215 +done
 20.1216 +
 20.1217 +lemma sublist_singleton [simp]:
 20.1218 +     "sublist([x], A) = (if 0 : A then [x] else [])"
 20.1219 +by (simp add: sublist_Cons)
 20.1220 +
 20.1221 +lemma sublist_upt_eq_take [rule_format, simp]:
 20.1222 +    "xs:list(A) ==> ALL n:nat. sublist(xs,n) = take(n,xs)"
 20.1223 +apply (erule list.induct, simp) 
 20.1224 +apply (clarify ); 
 20.1225 +apply (erule natE) 
 20.1226 +apply (simp_all add: nat_eq_Collect_lt Ord_mem_iff_lt sublist_Cons)
 20.1227 +done
 20.1228 +
 20.1229 +lemma sublist_Int_eq:
 20.1230 +     "xs : list(B) ==> sublist(xs, A \<inter> nat) = sublist(xs, A)"
 20.1231 +apply (erule list.induct)
 20.1232 +apply (simp_all add: sublist_Cons) 
 20.1233 +done
 20.1234 +
 20.1235 +text{*Repetition of a List Element*}
 20.1236 +
 20.1237 +consts   repeat :: "[i,i]=>i"
 20.1238 +primrec
 20.1239 +  "repeat(a,0) = []"
 20.1240 +
 20.1241 +  "repeat(a,succ(n)) = Cons(a,repeat(a,n))"
 20.1242 +
 20.1243 +lemma length_repeat: "n \<in> nat ==> length(repeat(a,n)) = n"
 20.1244 +by (induct_tac n, auto)
 20.1245 +
 20.1246 +lemma repeat_succ_app: "n \<in> nat ==> repeat(a,succ(n)) = repeat(a,n) @ [a]"
 20.1247 +apply (induct_tac n)
 20.1248 +apply (simp_all del: app_Cons add: app_Cons [symmetric])
 20.1249 +done
 20.1250 +
 20.1251 +lemma repeat_type [TC]: "[|a \<in> A; n \<in> nat|] ==> repeat(a,n) \<in> list(A)"
 20.1252 +by (induct_tac n, auto)
 20.1253 +
 20.1254 +end
    21.1 --- a/src/ZF/Main.thy	Mon Feb 11 15:19:17 2008 +0100
    21.2 +++ b/src/ZF/Main.thy	Mon Feb 11 15:40:21 2008 +0100
    21.3 @@ -1,79 +1,5 @@
    21.4 -(*$Id$*)
    21.5 -
    21.6 -header{*Theory Main: Everything Except AC*}
    21.7 -
    21.8 -theory Main imports List IntDiv CardinalArith begin
    21.9 -
   21.10 -(*The theory of "iterates" logically belongs to Nat, but can't go there because
   21.11 -  primrec isn't available into after Datatype.*)
   21.12 -
   21.13 -subsection{* Iteration of the function @{term F} *}
   21.14 -
   21.15 -consts  iterates :: "[i=>i,i,i] => i"   ("(_^_ '(_'))" [60,1000,1000] 60)
   21.16 -
   21.17 -primrec
   21.18 -    "F^0 (x) = x"
   21.19 -    "F^(succ(n)) (x) = F(F^n (x))"
   21.20 -
   21.21 -definition
   21.22 -  iterates_omega :: "[i=>i,i] => i"  where
   21.23 -    "iterates_omega(F,x) == \<Union>n\<in>nat. F^n (x)"
   21.24 -
   21.25 -notation (xsymbols)
   21.26 -  iterates_omega  ("(_^\<omega> '(_'))" [60,1000] 60)
   21.27 -notation (HTML output)
   21.28 -  iterates_omega  ("(_^\<omega> '(_'))" [60,1000] 60)
   21.29 -
   21.30 -lemma iterates_triv:
   21.31 -     "[| n\<in>nat;  F(x) = x |] ==> F^n (x) = x"  
   21.32 -by (induct n rule: nat_induct, simp_all)
   21.33 -
   21.34 -lemma iterates_type [TC]:
   21.35 -     "[| n:nat;  a: A; !!x. x:A ==> F(x) : A |] 
   21.36 -      ==> F^n (a) : A"  
   21.37 -by (induct n rule: nat_induct, simp_all)
   21.38 -
   21.39 -lemma iterates_omega_triv:
   21.40 -    "F(x) = x ==> F^\<omega> (x) = x" 
   21.41 -by (simp add: iterates_omega_def iterates_triv) 
   21.42 -
   21.43 -lemma Ord_iterates [simp]:
   21.44 -     "[| n\<in>nat;  !!i. Ord(i) ==> Ord(F(i));  Ord(x) |] 
   21.45 -      ==> Ord(F^n (x))"  
   21.46 -by (induct n rule: nat_induct, simp_all)
   21.47 -
   21.48 -lemma iterates_commute: "n \<in> nat ==> F(F^n (x)) = F^n (F(x))"
   21.49 -by (induct_tac n, simp_all)
   21.50 -
   21.51 -
   21.52 -subsection{* Transfinite Recursion *}
   21.53 -
   21.54 -text{*Transfinite recursion for definitions based on the 
   21.55 -    three cases of ordinals*}
   21.56 -
   21.57 -definition
   21.58 -  transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where
   21.59 -    "transrec3(k, a, b, c) ==                     
   21.60 -       transrec(k, \<lambda>x r.
   21.61 -         if x=0 then a
   21.62 -         else if Limit(x) then c(x, \<lambda>y\<in>x. r`y)
   21.63 -         else b(Arith.pred(x), r ` Arith.pred(x)))"
   21.64 -
   21.65 -lemma transrec3_0 [simp]: "transrec3(0,a,b,c) = a"
   21.66 -by (rule transrec3_def [THEN def_transrec, THEN trans], simp)
   21.67 -
   21.68 -lemma transrec3_succ [simp]:
   21.69 -     "transrec3(succ(i),a,b,c) = b(i, transrec3(i,a,b,c))"
   21.70 -by (rule transrec3_def [THEN def_transrec, THEN trans], simp)
   21.71 -
   21.72 -lemma transrec3_Limit:
   21.73 -     "Limit(i) ==> 
   21.74 -      transrec3(i,a,b,c) = c(i, \<lambda>j\<in>i. transrec3(j,a,b,c))"
   21.75 -by (rule transrec3_def [THEN def_transrec, THEN trans], force)
   21.76 -
   21.77 -
   21.78 -ML_setup {*
   21.79 -  change_simpset (fn ss => ss setmksimps (map mk_eq o Ord_atomize o gen_all));
   21.80 -*}
   21.81 +theory Main 
   21.82 +imports Main_ZF
   21.83 +begin
   21.84  
   21.85  end
    22.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    22.2 +++ b/src/ZF/Main_ZF.thy	Mon Feb 11 15:40:21 2008 +0100
    22.3 @@ -0,0 +1,79 @@
    22.4 +(*$Id$*)
    22.5 +
    22.6 +header{*Theory Main: Everything Except AC*}
    22.7 +
    22.8 +theory Main_ZF imports List_ZF IntDiv_ZF CardinalArith begin
    22.9 +
   22.10 +(*The theory of "iterates" logically belongs to Nat, but can't go there because
   22.11 +  primrec isn't available into after Datatype.*)
   22.12 +
   22.13 +subsection{* Iteration of the function @{term F} *}
   22.14 +
   22.15 +consts  iterates :: "[i=>i,i,i] => i"   ("(_^_ '(_'))" [60,1000,1000] 60)
   22.16 +
   22.17 +primrec
   22.18 +    "F^0 (x) = x"
   22.19 +    "F^(succ(n)) (x) = F(F^n (x))"
   22.20 +
   22.21 +definition
   22.22 +  iterates_omega :: "[i=>i,i] => i"  where
   22.23 +    "iterates_omega(F,x) == \<Union>n\<in>nat. F^n (x)"
   22.24 +
   22.25 +notation (xsymbols)
   22.26 +  iterates_omega  ("(_^\<omega> '(_'))" [60,1000] 60)
   22.27 +notation (HTML output)
   22.28 +  iterates_omega  ("(_^\<omega> '(_'))" [60,1000] 60)
   22.29 +
   22.30 +lemma iterates_triv:
   22.31 +     "[| n\<in>nat;  F(x) = x |] ==> F^n (x) = x"  
   22.32 +by (induct n rule: nat_induct, simp_all)
   22.33 +
   22.34 +lemma iterates_type [TC]:
   22.35 +     "[| n:nat;  a: A; !!x. x:A ==> F(x) : A |] 
   22.36 +      ==> F^n (a) : A"  
   22.37 +by (induct n rule: nat_induct, simp_all)
   22.38 +
   22.39 +lemma iterates_omega_triv:
   22.40 +    "F(x) = x ==> F^\<omega> (x) = x" 
   22.41 +by (simp add: iterates_omega_def iterates_triv) 
   22.42 +
   22.43 +lemma Ord_iterates [simp]:
   22.44 +     "[| n\<in>nat;  !!i. Ord(i) ==> Ord(F(i));  Ord(x) |] 
   22.45 +      ==> Ord(F^n (x))"  
   22.46 +by (induct n rule: nat_induct, simp_all)
   22.47 +
   22.48 +lemma iterates_commute: "n \<in> nat ==> F(F^n (x)) = F^n (F(x))"
   22.49 +by (induct_tac n, simp_all)
   22.50 +
   22.51 +
   22.52 +subsection{* Transfinite Recursion *}
   22.53 +
   22.54 +text{*Transfinite recursion for definitions based on the 
   22.55 +    three cases of ordinals*}
   22.56 +
   22.57 +definition
   22.58 +  transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where
   22.59 +    "transrec3(k, a, b, c) ==                     
   22.60 +       transrec(k, \<lambda>x r.
   22.61 +         if x=0 then a
   22.62 +         else if Limit(x) then c(x, \<lambda>y\<in>x. r`y)
   22.63 +         else b(Arith.pred(x), r ` Arith.pred(x)))"
   22.64 +
   22.65 +lemma transrec3_0 [simp]: "transrec3(0,a,b,c) = a"
   22.66 +by (rule transrec3_def [THEN def_transrec, THEN trans], simp)
   22.67 +
   22.68 +lemma transrec3_succ [simp]:
   22.69 +     "transrec3(succ(i),a,b,c) = b(i, transrec3(i,a,b,c))"
   22.70 +by (rule transrec3_def [THEN def_transrec, THEN trans], simp)
   22.71 +
   22.72 +lemma transrec3_Limit:
   22.73 +     "Limit(i) ==> 
   22.74 +      transrec3(i,a,b,c) = c(i, \<lambda>j\<in>i. transrec3(j,a,b,c))"
   22.75 +by (rule transrec3_def [THEN def_transrec, THEN trans], force)
   22.76 +
   22.77 +
   22.78 +ML_setup {*
   22.79 +  change_simpset (fn ss => ss setmksimps (map mk_eq o Ord_atomize o gen_all));
   22.80 +*}
   22.81 +
   22.82 +end
    23.1 --- a/src/ZF/Main_ZFC.thy	Mon Feb 11 15:19:17 2008 +0100
    23.2 +++ b/src/ZF/Main_ZFC.thy	Mon Feb 11 15:40:21 2008 +0100
    23.3 @@ -1,3 +1,3 @@
    23.4 -theory Main_ZFC imports Main InfDatatype begin
    23.5 +theory Main_ZFC imports Main_ZF InfDatatype begin
    23.6  
    23.7  end
    24.1 --- a/src/ZF/Nat.thy	Mon Feb 11 15:19:17 2008 +0100
    24.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    24.3 @@ -1,302 +0,0 @@
    24.4 -(*  Title:      ZF/Nat.thy
    24.5 -    ID:         $Id$
    24.6 -    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    24.7 -    Copyright   1994  University of Cambridge
    24.8 -
    24.9 -*)
   24.10 -
   24.11 -header{*The Natural numbers As a Least Fixed Point*}
   24.12 -
   24.13 -theory Nat imports OrdQuant Bool begin
   24.14 -
   24.15 -definition
   24.16 -  nat :: i  where
   24.17 -    "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})"
   24.18 -
   24.19 -definition
   24.20 -  quasinat :: "i => o"  where
   24.21 -    "quasinat(n) == n=0 | (\<exists>m. n = succ(m))"
   24.22 -
   24.23 -definition
   24.24 -  (*Has an unconditional succ case, which is used in "recursor" below.*)
   24.25 -  nat_case :: "[i, i=>i, i]=>i"  where
   24.26 -    "nat_case(a,b,k) == THE y. k=0 & y=a | (EX x. k=succ(x) & y=b(x))"
   24.27 -
   24.28 -definition
   24.29 -  nat_rec :: "[i, i, [i,i]=>i]=>i"  where
   24.30 -    "nat_rec(k,a,b) ==   
   24.31 -          wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))"
   24.32 -
   24.33 -  (*Internalized relations on the naturals*)
   24.34 -  
   24.35 -definition
   24.36 -  Le :: i  where
   24.37 -    "Le == {<x,y>:nat*nat. x le y}"
   24.38 -
   24.39 -definition
   24.40 -  Lt :: i  where
   24.41 -    "Lt == {<x, y>:nat*nat. x < y}"
   24.42 -  
   24.43 -definition
   24.44 -  Ge :: i  where
   24.45 -    "Ge == {<x,y>:nat*nat. y le x}"
   24.46 -
   24.47 -definition
   24.48 -  Gt :: i  where
   24.49 -    "Gt == {<x,y>:nat*nat. y < x}"
   24.50 -
   24.51 -definition
   24.52 -  greater_than :: "i=>i"  where
   24.53 -    "greater_than(n) == {i:nat. n < i}"
   24.54 -
   24.55 -text{*No need for a less-than operator: a natural number is its list of
   24.56 -predecessors!*}
   24.57 -
   24.58 -
   24.59 -lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})"
   24.60 -apply (rule bnd_monoI)
   24.61 -apply (cut_tac infinity, blast, blast) 
   24.62 -done
   24.63 -
   24.64 -(* nat = {0} Un {succ(x). x:nat} *)
   24.65 -lemmas nat_unfold = nat_bnd_mono [THEN nat_def [THEN def_lfp_unfold], standard]
   24.66 -
   24.67 -(** Type checking of 0 and successor **)
   24.68 -
   24.69 -lemma nat_0I [iff,TC]: "0 : nat"
   24.70 -apply (subst nat_unfold)
   24.71 -apply (rule singletonI [THEN UnI1])
   24.72 -done
   24.73 -
   24.74 -lemma nat_succI [intro!,TC]: "n : nat ==> succ(n) : nat"
   24.75 -apply (subst nat_unfold)
   24.76 -apply (erule RepFunI [THEN UnI2])
   24.77 -done
   24.78 -
   24.79 -lemma nat_1I [iff,TC]: "1 : nat"
   24.80 -by (rule nat_0I [THEN nat_succI])
   24.81 -
   24.82 -lemma nat_2I [iff,TC]: "2 : nat"
   24.83 -by (rule nat_1I [THEN nat_succI])
   24.84 -
   24.85 -lemma bool_subset_nat: "bool <= nat"
   24.86 -by (blast elim!: boolE)
   24.87 -
   24.88 -lemmas bool_into_nat = bool_subset_nat [THEN subsetD, standard]
   24.89 -
   24.90 -
   24.91 -subsection{*Injectivity Properties and Induction*}
   24.92 -
   24.93 -(*Mathematical induction*)
   24.94 -lemma nat_induct [case_names 0 succ, induct set: nat]:
   24.95 -    "[| n: nat;  P(0);  !!x. [| x: nat;  P(x) |] ==> P(succ(x)) |] ==> P(n)"
   24.96 -by (erule def_induct [OF nat_def nat_bnd_mono], blast)
   24.97 -
   24.98 -lemma natE:
   24.99 -    "[| n: nat;  n=0 ==> P;  !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P"
  24.100 -by (erule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE], auto) 
  24.101 -
  24.102 -lemma nat_into_Ord [simp]: "n: nat ==> Ord(n)"
  24.103 -by (erule nat_induct, auto)
  24.104 -
  24.105 -(* i: nat ==> 0 le i; same thing as 0<succ(i)  *)
  24.106 -lemmas nat_0_le = nat_into_Ord [THEN Ord_0_le, standard]
  24.107 -
  24.108 -(* i: nat ==> i le i; same thing as i<succ(i)  *)
  24.109 -lemmas nat_le_refl = nat_into_Ord [THEN le_refl, standard]
  24.110 -
  24.111 -lemma Ord_nat [iff]: "Ord(nat)"
  24.112 -apply (rule OrdI)
  24.113 -apply (erule_tac [2] nat_into_Ord [THEN Ord_is_Transset])
  24.114 -apply (unfold Transset_def)
  24.115 -apply (rule ballI)
  24.116 -apply (erule nat_induct, auto) 
  24.117 -done
  24.118 -
  24.119 -lemma Limit_nat [iff]: "Limit(nat)"
  24.120 -apply (unfold Limit_def)
  24.121 -apply (safe intro!: ltI Ord_nat)
  24.122 -apply (erule ltD)
  24.123 -done
  24.124 -
  24.125 -lemma naturals_not_limit: "a \<in> nat ==> ~ Limit(a)"
  24.126 -by (induct a rule: nat_induct, auto)
  24.127 -
  24.128 -lemma succ_natD: "succ(i): nat ==> i: nat"
  24.129 -by (rule Ord_trans [OF succI1], auto)
  24.130 -
  24.131 -lemma nat_succ_iff [iff]: "succ(n): nat <-> n: nat"
  24.132 -by (blast dest!: succ_natD)
  24.133 -
  24.134 -lemma nat_le_Limit: "Limit(i) ==> nat le i"
  24.135 -apply (rule subset_imp_le)
  24.136 -apply (simp_all add: Limit_is_Ord) 
  24.137 -apply (rule subsetI)
  24.138 -apply (erule nat_induct)
  24.139 - apply (erule Limit_has_0 [THEN ltD]) 
  24.140 -apply (blast intro: Limit_has_succ [THEN ltD] ltI Limit_is_Ord)
  24.141 -done
  24.142 -
  24.143 -(* [| succ(i): k;  k: nat |] ==> i: k *)
  24.144 -lemmas succ_in_naturalD = Ord_trans [OF succI1 _ nat_into_Ord]
  24.145 -
  24.146 -lemma lt_nat_in_nat: "[| m<n;  n: nat |] ==> m: nat"
  24.147 -apply (erule ltE)
  24.148 -apply (erule Ord_trans, assumption, simp) 
  24.149 -done
  24.150 -
  24.151 -lemma le_in_nat: "[| m le n; n:nat |] ==> m:nat"
  24.152 -by (blast dest!: lt_nat_in_nat)
  24.153 -
  24.154 -
  24.155 -subsection{*Variations on Mathematical Induction*}
  24.156 -
  24.157 -(*complete induction*)
  24.158 -
  24.159 -lemmas complete_induct = Ord_induct [OF _ Ord_nat, case_names less, consumes 1]
  24.160 -
  24.161 -lemmas complete_induct_rule =  
  24.162 -	complete_induct [rule_format, case_names less, consumes 1]
  24.163 -
  24.164 -
  24.165 -lemma nat_induct_from_lemma [rule_format]: 
  24.166 -    "[| n: nat;  m: nat;   
  24.167 -        !!x. [| x: nat;  m le x;  P(x) |] ==> P(succ(x)) |] 
  24.168 -     ==> m le n --> P(m) --> P(n)"
  24.169 -apply (erule nat_induct) 
  24.170 -apply (simp_all add: distrib_simps le0_iff le_succ_iff)
  24.171 -done
  24.172 -
  24.173 -(*Induction starting from m rather than 0*)
  24.174 -lemma nat_induct_from: 
  24.175 -    "[| m le n;  m: nat;  n: nat;   
  24.176 -        P(m);   
  24.177 -        !!x. [| x: nat;  m le x;  P(x) |] ==> P(succ(x)) |]
  24.178 -     ==> P(n)"
  24.179 -apply (blast intro: nat_induct_from_lemma)
  24.180 -done
  24.181 -
  24.182 -(*Induction suitable for subtraction and less-than*)
  24.183 -lemma diff_induct [case_names 0 0_succ succ_succ, consumes 2]:
  24.184 -    "[| m: nat;  n: nat;   
  24.185 -        !!x. x: nat ==> P(x,0);   
  24.186 -        !!y. y: nat ==> P(0,succ(y));   
  24.187 -        !!x y. [| x: nat;  y: nat;  P(x,y) |] ==> P(succ(x),succ(y)) |]
  24.188 -     ==> P(m,n)"
  24.189 -apply (erule_tac x = m in rev_bspec)
  24.190 -apply (erule nat_induct, simp) 
  24.191 -apply (rule ballI)
  24.192 -apply (rename_tac i j)
  24.193 -apply (erule_tac n=j in nat_induct, auto)  
  24.194 -done
  24.195 -
  24.196 -
  24.197 -(** Induction principle analogous to trancl_induct **)
  24.198 -
  24.199 -lemma succ_lt_induct_lemma [rule_format]:
  24.200 -     "m: nat ==> P(m,succ(m)) --> (ALL x: nat. P(m,x) --> P(m,succ(x))) -->  
  24.201 -                 (ALL n:nat. m<n --> P(m,n))"
  24.202 -apply (erule nat_induct)
  24.203 - apply (intro impI, rule nat_induct [THEN ballI])
  24.204 -   prefer 4 apply (intro impI, rule nat_induct [THEN ballI])
  24.205 -apply (auto simp add: le_iff) 
  24.206 -done
  24.207 -
  24.208 -lemma succ_lt_induct:
  24.209 -    "[| m<n;  n: nat;                                    
  24.210 -        P(m,succ(m));                                    
  24.211 -        !!x. [| x: nat;  P(m,x) |] ==> P(m,succ(x)) |]
  24.212 -     ==> P(m,n)"
  24.213 -by (blast intro: succ_lt_induct_lemma lt_nat_in_nat) 
  24.214 -
  24.215 -subsection{*quasinat: to allow a case-split rule for @{term nat_case}*}
  24.216 -
  24.217 -text{*True if the argument is zero or any successor*}
  24.218 -lemma [iff]: "quasinat(0)"
  24.219 -by (simp add: quasinat_def)
  24.220 -
  24.221 -lemma [iff]: "quasinat(succ(x))"
  24.222 -by (simp add: quasinat_def)
  24.223 -
  24.224 -lemma nat_imp_quasinat: "n \<in> nat ==> quasinat(n)"
  24.225 -by (erule natE, simp_all)
  24.226 -
  24.227 -lemma non_nat_case: "~ quasinat(x) ==> nat_case(a,b,x) = 0" 
  24.228 -by (simp add: quasinat_def nat_case_def) 
  24.229 -
  24.230 -lemma nat_cases_disj: "k=0 | (\<exists>y. k = succ(y)) | ~ quasinat(k)"
  24.231 -apply (case_tac "k=0", simp) 
  24.232 -apply (case_tac "\<exists>m. k = succ(m)") 
  24.233 -apply (simp_all add: quasinat_def) 
  24.234 -done
  24.235 -
  24.236 -lemma nat_cases:
  24.237 -     "[|k=0 ==> P;  !!y. k = succ(y) ==> P; ~ quasinat(k) ==> P|] ==> P"
  24.238 -by (insert nat_cases_disj [of k], blast) 
  24.239 -
  24.240 -(** nat_case **)
  24.241 -
  24.242 -lemma nat_case_0 [simp]: "nat_case(a,b,0) = a"
  24.243 -by (simp add: nat_case_def)
  24.244 -
  24.245 -lemma nat_case_succ [simp]: "nat_case(a,b,succ(n)) = b(n)" 
  24.246 -by (simp add: nat_case_def)
  24.247 -
  24.248 -lemma nat_case_type [TC]:
  24.249 -    "[| n: nat;  a: C(0);  !!m. m: nat ==> b(m): C(succ(m)) |] 
  24.250 -     ==> nat_case(a,b,n) : C(n)";
  24.251 -by (erule nat_induct, auto) 
  24.252 -
  24.253 -lemma split_nat_case:
  24.254 -  "P(nat_case(a,b,k)) <-> 
  24.255 -   ((k=0 --> P(a)) & (\<forall>x. k=succ(x) --> P(b(x))) & (~ quasinat(k) \<longrightarrow> P(0)))"
  24.256 -apply (rule nat_cases [of k]) 
  24.257 -apply (auto simp add: non_nat_case)
  24.258 -done
  24.259 -
  24.260 -
  24.261 -subsection{*Recursion on the Natural Numbers*}
  24.262 -
  24.263 -(** nat_rec is used to define eclose and transrec, then becomes obsolete.
  24.264 -    The operator rec, from arith.thy, has fewer typing conditions **)
  24.265 -
  24.266 -lemma nat_rec_0: "nat_rec(0,a,b) = a"
  24.267 -apply (rule nat_rec_def [THEN def_wfrec, THEN trans])
  24.268 - apply (rule wf_Memrel) 
  24.269 -apply (rule nat_case_0)
  24.270 -done
  24.271 -
  24.272 -lemma nat_rec_succ: "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"
  24.273 -apply (rule nat_rec_def [THEN def_wfrec, THEN trans])
  24.274 - apply (rule wf_Memrel) 
  24.275 -apply (simp add: vimage_singleton_iff)
  24.276 -done
  24.277 -
  24.278 -(** The union of two natural numbers is a natural number -- their maximum **)
  24.279 -
  24.280 -lemma Un_nat_type [TC]: "[| i: nat; j: nat |] ==> i Un j: nat"
  24.281 -apply (rule Un_least_lt [THEN ltD])
  24.282 -apply (simp_all add: lt_def) 
  24.283 -done
  24.284 -
  24.285 -lemma Int_nat_type [TC]: "[| i: nat; j: nat |] ==> i Int j: nat"
  24.286 -apply (rule Int_greatest_lt [THEN ltD])
  24.287 -apply (simp_all add: lt_def) 
  24.288 -done
  24.289 -
  24.290 -(*needed to simplify unions over nat*)
  24.291 -lemma nat_nonempty [simp]: "nat ~= 0"
  24.292 -by blast
  24.293 -
  24.294 -text{*A natural number is the set of its predecessors*}
  24.295 -lemma nat_eq_Collect_lt: "i \<in> nat ==> {j\<in>nat. j<i} = i"
  24.296 -apply (rule equalityI)
  24.297 -apply (blast dest: ltD)  
  24.298 -apply (auto simp add: Ord_mem_iff_lt)
  24.299 -apply (blast intro: lt_trans) 
  24.300 -done
  24.301 -
  24.302 -lemma Le_iff [iff]: "<x,y> : Le <-> x le y & x : nat & y : nat"
  24.303 -by (force simp add: Le_def)
  24.304 -
  24.305 -end
    25.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    25.2 +++ b/src/ZF/Nat_ZF.thy	Mon Feb 11 15:40:21 2008 +0100
    25.3 @@ -0,0 +1,302 @@
    25.4 +(*  Title:      ZF/Nat.thy
    25.5 +    ID:         $Id$
    25.6 +    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    25.7 +    Copyright   1994  University of Cambridge
    25.8 +
    25.9 +*)
   25.10 +
   25.11 +header{*The Natural numbers As a Least Fixed Point*}
   25.12 +
   25.13 +theory Nat_ZF imports OrdQuant Bool begin
   25.14 +
   25.15 +definition
   25.16 +  nat :: i  where
   25.17 +    "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})"
   25.18 +
   25.19 +definition
   25.20 +  quasinat :: "i => o"  where
   25.21 +    "quasinat(n) == n=0 | (\<exists>m. n = succ(m))"
   25.22 +
   25.23 +definition
   25.24 +  (*Has an unconditional succ case, which is used in "recursor" below.*)
   25.25 +  nat_case :: "[i, i=>i, i]=>i"  where
   25.26 +    "nat_case(a,b,k) == THE y. k=0 & y=a | (EX x. k=succ(x) & y=b(x))"
   25.27 +
   25.28 +definition
   25.29 +  nat_rec :: "[i, i, [i,i]=>i]=>i"  where
   25.30 +    "nat_rec(k,a,b) ==   
   25.31 +          wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))"
   25.32 +
   25.33 +  (*Internalized relations on the naturals*)
   25.34 +  
   25.35 +definition
   25.36 +  Le :: i  where
   25.37 +    "Le == {<x,y>:nat*nat. x le y}"
   25.38 +
   25.39 +definition
   25.40 +  Lt :: i  where
   25.41 +    "Lt == {<x, y>:nat*nat. x < y}"
   25.42 +  
   25.43 +definition
   25.44 +  Ge :: i  where
   25.45 +    "Ge == {<x,y>:nat*nat. y le x}"
   25.46 +
   25.47 +definition
   25.48 +  Gt :: i  where
   25.49 +    "Gt == {<x,y>:nat*nat. y < x}"
   25.50 +
   25.51 +definition
   25.52 +  greater_than :: "i=>i"  where
   25.53 +    "greater_than(n) == {i:nat. n < i}"
   25.54 +
   25.55 +text{*No need for a less-than operator: a natural number is its list of
   25.56 +predecessors!*}
   25.57 +
   25.58 +
   25.59 +lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})"
   25.60 +apply (rule bnd_monoI)
   25.61 +apply (cut_tac infinity, blast, blast) 
   25.62 +done
   25.63 +
   25.64 +(* nat = {0} Un {succ(x). x:nat} *)
   25.65 +lemmas nat_unfold = nat_bnd_mono [THEN nat_def [THEN def_lfp_unfold], standard]
   25.66 +
   25.67 +(** Type checking of 0 and successor **)
   25.68 +
   25.69 +lemma nat_0I [iff,TC]: "0 : nat"
   25.70 +apply (subst nat_unfold)
   25.71 +apply (rule singletonI [THEN UnI1])
   25.72 +done
   25.73 +
   25.74 +lemma nat_succI [intro!,TC]: "n : nat ==> succ(n) : nat"
   25.75 +apply (subst nat_unfold)
   25.76 +apply (erule RepFunI [THEN UnI2])
   25.77 +done
   25.78 +
   25.79 +lemma nat_1I [iff,TC]: "1 : nat"
   25.80 +by (rule nat_0I [THEN nat_succI])
   25.81 +
   25.82 +lemma nat_2I [iff,TC]: "2 : nat"
   25.83 +by (rule nat_1I [THEN nat_succI])
   25.84 +
   25.85 +lemma bool_subset_nat: "bool <= nat"
   25.86 +by (blast elim!: boolE)
   25.87 +
   25.88 +lemmas bool_into_nat = bool_subset_nat [THEN subsetD, standard]
   25.89 +
   25.90 +
   25.91 +subsection{*Injectivity Properties and Induction*}
   25.92 +
   25.93 +(*Mathematical induction*)
   25.94 +lemma nat_induct [case_names 0 succ, induct set: nat]:
   25.95 +    "[| n: nat;  P(0);  !!x. [| x: nat;  P(x) |] ==> P(succ(x)) |] ==> P(n)"
   25.96 +by (erule def_induct [OF nat_def nat_bnd_mono], blast)
   25.97 +
   25.98 +lemma natE:
   25.99 +    "[| n: nat;  n=0 ==> P;  !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P"
  25.100 +by (erule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE], auto) 
  25.101 +
  25.102 +lemma nat_into_Ord [simp]: "n: nat ==> Ord(n)"
  25.103 +by (erule nat_induct, auto)
  25.104 +
  25.105 +(* i: nat ==> 0 le i; same thing as 0<succ(i)  *)
  25.106 +lemmas nat_0_le = nat_into_Ord [THEN Ord_0_le, standard]
  25.107 +
  25.108 +(* i: nat ==> i le i; same thing as i<succ(i)  *)
  25.109 +lemmas nat_le_refl = nat_into_Ord [THEN le_refl, standard]
  25.110 +
  25.111 +lemma Ord_nat [iff]: "Ord(nat)"
  25.112 +apply (rule OrdI)
  25.113 +apply (erule_tac [2] nat_into_Ord [THEN Ord_is_Transset])
  25.114 +apply (unfold Transset_def)
  25.115 +apply (rule ballI)
  25.116 +apply (erule nat_induct, auto) 
  25.117 +done
  25.118 +
  25.119 +lemma Limit_nat [iff]: "Limit(nat)"
  25.120 +apply (unfold Limit_def)
  25.121 +apply (safe intro!: ltI Ord_nat)
  25.122 +apply (erule ltD)
  25.123 +done
  25.124 +
  25.125 +lemma naturals_not_limit: "a \<in> nat ==> ~ Limit(a)"
  25.126 +by (induct a rule: nat_induct, auto)
  25.127 +
  25.128 +lemma succ_natD: "succ(i): nat ==> i: nat"
  25.129 +by (rule Ord_trans [OF succI1], auto)
  25.130 +
  25.131 +lemma nat_succ_iff [iff]: "succ(n): nat <-> n: nat"
  25.132 +by (blast dest!: succ_natD)
  25.133 +
  25.134 +lemma nat_le_Limit: "Limit(i) ==> nat le i"
  25.135 +apply (rule subset_imp_le)
  25.136 +apply (simp_all add: Limit_is_Ord) 
  25.137 +apply (rule subsetI)
  25.138 +apply (erule nat_induct)
  25.139 + apply (erule Limit_has_0 [THEN ltD]) 
  25.140 +apply (blast intro: Limit_has_succ [THEN ltD] ltI Limit_is_Ord)
  25.141 +done
  25.142 +
  25.143 +(* [| succ(i): k;  k: nat |] ==> i: k *)
  25.144 +lemmas succ_in_naturalD = Ord_trans [OF succI1 _ nat_into_Ord]
  25.145 +
  25.146 +lemma lt_nat_in_nat: "[| m<n;  n: nat |] ==> m: nat"
  25.147 +apply (erule ltE)
  25.148 +apply (erule Ord_trans, assumption, simp) 
  25.149 +done
  25.150 +
  25.151 +lemma le_in_nat: "[| m le n; n:nat |] ==> m:nat"
  25.152 +by (blast dest!: lt_nat_in_nat)
  25.153 +
  25.154 +
  25.155 +subsection{*Variations on Mathematical Induction*}
  25.156 +
  25.157 +(*complete induction*)
  25.158 +
  25.159 +lemmas complete_induct = Ord_induct [OF _ Ord_nat, case_names less, consumes 1]
  25.160 +
  25.161 +lemmas complete_induct_rule =  
  25.162 +	complete_induct [rule_format, case_names less, consumes 1]
  25.163 +
  25.164 +
  25.165 +lemma nat_induct_from_lemma [rule_format]: 
  25.166 +    "[| n: nat;  m: nat;   
  25.167 +        !!x. [| x: nat;  m le x;  P(x) |] ==> P(succ(x)) |] 
  25.168 +     ==> m le n --> P(m) --> P(n)"
  25.169 +apply (erule nat_induct) 
  25.170 +apply (simp_all add: distrib_simps le0_iff le_succ_iff)
  25.171 +done
  25.172 +
  25.173 +(*Induction starting from m rather than 0*)
  25.174 +lemma nat_induct_from: 
  25.175 +    "[| m le n;  m: nat;  n: nat;   
  25.176 +        P(m);   
  25.177 +        !!x. [| x: nat;  m le x;  P(x) |] ==> P(succ(x)) |]
  25.178 +     ==> P(n)"
  25.179 +apply (blast intro: nat_induct_from_lemma)
  25.180 +done
  25.181 +
  25.182 +(*Induction suitable for subtraction and less-than*)
  25.183 +lemma diff_induct [case_names 0 0_succ succ_succ, consumes 2]:
  25.184 +    "[| m: nat;  n: nat;   
  25.185 +        !!x. x: nat ==> P(x,0);   
  25.186 +        !!y. y: nat ==> P(0,succ(y));   
  25.187 +        !!x y. [| x: nat;  y: nat;  P(x,y) |] ==> P(succ(x),succ(y)) |]
  25.188 +     ==> P(m,n)"
  25.189 +apply (erule_tac x = m in rev_bspec)
  25.190 +apply (erule nat_induct, simp) 
  25.191 +apply (rule ballI)
  25.192 +apply (rename_tac i j)
  25.193 +apply (erule_tac n=j in nat_induct, auto)  
  25.194 +done
  25.195 +
  25.196 +
  25.197 +(** Induction principle analogous to trancl_induct **)
  25.198 +
  25.199 +lemma succ_lt_induct_lemma [rule_format]:
  25.200 +     "m: nat ==> P(m,succ(m)) --> (ALL x: nat. P(m,x) --> P(m,succ(x))) -->  
  25.201 +                 (ALL n:nat. m<n --> P(m,n))"
  25.202 +apply (erule nat_induct)
  25.203 + apply (intro impI, rule nat_induct [THEN ballI])
  25.204 +   prefer 4 apply (intro impI, rule nat_induct [THEN ballI])
  25.205 +apply (auto simp add: le_iff) 
  25.206 +done
  25.207 +
  25.208 +lemma succ_lt_induct:
  25.209 +    "[| m<n;  n: nat;                                    
  25.210 +        P(m,succ(m));                                    
  25.211 +        !!x. [| x: nat;  P(m,x) |] ==> P(m,succ(x)) |]
  25.212 +     ==> P(m,n)"
  25.213 +by (blast intro: succ_lt_induct_lemma lt_nat_in_nat) 
  25.214 +
  25.215 +subsection{*quasinat: to allow a case-split rule for @{term nat_case}*}
  25.216 +
  25.217 +text{*True if the argument is zero or any successor*}
  25.218 +lemma [iff]: "quasinat(0)"
  25.219 +by (simp add: quasinat_def)
  25.220 +
  25.221 +lemma [iff]: "quasinat(succ(x))"
  25.222 +by (simp add: quasinat_def)
  25.223 +
  25.224 +lemma nat_imp_quasinat: "n \<in> nat ==> quasinat(n)"
  25.225 +by (erule natE, simp_all)
  25.226 +
  25.227 +lemma non_nat_case: "~ quasinat(x) ==> nat_case(a,b,x) = 0" 
  25.228 +by (simp add: quasinat_def nat_case_def) 
  25.229 +
  25.230 +lemma nat_cases_disj: "k=0 | (\<exists>y. k = succ(y)) | ~ quasinat(k)"
  25.231 +apply (case_tac "k=0", simp) 
  25.232 +apply (case_tac "\<exists>m. k = succ(m)") 
  25.233 +apply (simp_all add: quasinat_def) 
  25.234 +done
  25.235 +
  25.236 +lemma nat_cases:
  25.237 +     "[|k=0 ==> P;  !!y. k = succ(y) ==> P; ~ quasinat(k) ==> P|] ==> P"
  25.238 +by (insert nat_cases_disj [of k], blast) 
  25.239 +
  25.240 +(** nat_case **)
  25.241 +
  25.242 +lemma nat_case_0 [simp]: "nat_case(a,b,0) = a"
  25.243 +by (simp add: nat_case_def)
  25.244 +
  25.245 +lemma nat_case_succ [simp]: "nat_case(a,b,succ(n)) = b(n)" 
  25.246 +by (simp add: nat_case_def)
  25.247 +
  25.248 +lemma nat_case_type [TC]:
  25.249 +    "[| n: nat;  a: C(0);  !!m. m: nat ==> b(m): C(succ(m)) |] 
  25.250 +     ==> nat_case(a,b,n) : C(n)";
  25.251 +by (erule nat_induct, auto) 
  25.252 +
  25.253 +lemma split_nat_case:
  25.254 +  "P(nat_case(a,b,k)) <-> 
  25.255 +   ((k=0 --> P(a)) & (\<forall>x. k=succ(x) --> P(b(x))) & (~ quasinat(k) \<longrightarrow> P(0)))"
  25.256 +apply (rule nat_cases [of k]) 
  25.257 +apply (auto simp add: non_nat_case)
  25.258 +done
  25.259 +
  25.260 +
  25.261 +subsection{*Recursion on the Natural Numbers*}
  25.262 +
  25.263 +(** nat_rec is used to define eclose and transrec, then becomes obsolete.
  25.264 +    The operator rec, from arith.thy, has fewer typing conditions **)
  25.265 +
  25.266 +lemma nat_rec_0: "nat_rec(0,a,b) = a"
  25.267 +apply (rule nat_rec_def [THEN def_wfrec, THEN trans])
  25.268 + apply (rule wf_Memrel) 
  25.269 +apply (rule nat_case_0)
  25.270 +done
  25.271 +
  25.272 +lemma nat_rec_succ: "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"
  25.273 +apply (rule nat_rec_def [THEN def_wfrec, THEN trans])
  25.274 + apply (rule wf_Memrel) 
  25.275 +apply (simp add: vimage_singleton_iff)
  25.276 +done
  25.277 +
  25.278 +(** The union of two natural numbers is a natural number -- their maximum **)
  25.279 +
  25.280 +lemma Un_nat_type [TC]: "[| i: nat; j: nat |] ==> i Un j: nat"
  25.281 +apply (rule Un_least_lt [THEN ltD])
  25.282 +apply (simp_all add: lt_def) 
  25.283 +done
  25.284 +
  25.285 +lemma Int_nat_type [TC]: "[| i: nat; j: nat |] ==> i Int j: nat"
  25.286 +apply (rule Int_greatest_lt [THEN ltD])
  25.287 +apply (simp_all add: lt_def) 
  25.288 +done
  25.289 +
  25.290 +(*needed to simplify unions over nat*)
  25.291 +lemma nat_nonempty [simp]: "nat ~= 0"
  25.292 +by blast
  25.293 +
  25.294 +text{*A natural number is the set of its predecessors*}
  25.295 +lemma nat_eq_Collect_lt: "i \<in> nat ==> {j\<in>nat. j<i} = i"
  25.296 +apply (rule equalityI)
  25.297 +apply (blast dest: ltD)  
  25.298 +apply (auto simp add: Ord_mem_iff_lt)
  25.299 +apply (blast intro: lt_trans) 
  25.300 +done
  25.301 +
  25.302 +lemma Le_iff [iff]: "<x,y> : Le <-> x le y & x : nat & y : nat"
  25.303 +by (force simp add: Le_def)
  25.304 +
  25.305 +end
    26.1 --- a/src/ZF/OrderType.thy	Mon Feb 11 15:19:17 2008 +0100
    26.2 +++ b/src/ZF/OrderType.thy	Mon Feb 11 15:40:21 2008 +0100
    26.3 @@ -7,7 +7,7 @@
    26.4  
    26.5  header{*Order Types and Ordinal Arithmetic*}
    26.6  
    26.7 -theory OrderType imports OrderArith OrdQuant Nat begin
    26.8 +theory OrderType imports OrderArith OrdQuant Nat_ZF begin
    26.9  
   26.10  text{*The order type of a well-ordering is the least ordinal isomorphic to it.
   26.11  Ordinal arithmetic is traditionally defined in terms of order types, as it is
    27.1 --- a/src/ZF/ROOT.ML	Mon Feb 11 15:19:17 2008 +0100
    27.2 +++ b/src/ZF/ROOT.ML	Mon Feb 11 15:40:21 2008 +0100
    27.3 @@ -8,5 +8,6 @@
    27.4  Paulson.
    27.5  *)
    27.6  
    27.7 +use_thy "Main";
    27.8  use_thy "Main_ZFC";
    27.9  
    28.1 --- a/src/ZF/Tools/datatype_package.ML	Mon Feb 11 15:19:17 2008 +0100
    28.2 +++ b/src/ZF/Tools/datatype_package.ML	Mon Feb 11 15:40:21 2008 +0100
    28.3 @@ -49,7 +49,7 @@
    28.4  fun add_datatype_i (dom_sum, rec_tms) con_ty_lists (monos, type_intrs, type_elims) thy =
    28.5   let
    28.6    val dummy = (*has essential ancestors?*)
    28.7 -    Theory.requires thy "Datatype" "(co)datatype definitions";
    28.8 +    Theory.requires thy "Datatype_ZF" "(co)datatype definitions";
    28.9  
   28.10    val rec_hds = map head_of rec_tms;
   28.11  
    29.1 --- a/src/ZF/Tools/inductive_package.ML	Mon Feb 11 15:19:17 2008 +0100
    29.2 +++ b/src/ZF/Tools/inductive_package.ML	Mon Feb 11 15:40:21 2008 +0100
    29.3 @@ -64,7 +64,7 @@
    29.4  fun add_inductive_i verbose (rec_tms, dom_sum)
    29.5    intr_specs (monos, con_defs, type_intrs, type_elims) thy =
    29.6  let
    29.7 -  val _ = Theory.requires thy "Inductive" "(co)inductive definitions";
    29.8 +  val _ = Theory.requires thy "Inductive_ZF" "(co)inductive definitions";
    29.9  
   29.10    val (intr_names, intr_tms) = split_list (map fst intr_specs);
   29.11    val case_names = RuleCases.case_names intr_names;
    30.1 --- a/src/ZF/UNITY/GenPrefix.thy	Mon Feb 11 15:19:17 2008 +0100
    30.2 +++ b/src/ZF/UNITY/GenPrefix.thy	Mon Feb 11 15:40:21 2008 +0100
    30.3 @@ -12,7 +12,7 @@
    30.4  header{*Charpentier's Generalized Prefix Relation*}
    30.5  
    30.6  theory GenPrefix
    30.7 -imports Main
    30.8 +imports Main_ZF
    30.9  begin
   30.10  
   30.11  definition (*really belongs in ZF/Trancl*)
    31.1 --- a/src/ZF/UNITY/Monotonicity.thy	Mon Feb 11 15:19:17 2008 +0100
    31.2 +++ b/src/ZF/UNITY/Monotonicity.thy	Mon Feb 11 15:40:21 2008 +0100
    31.3 @@ -69,7 +69,7 @@
    31.4  lemma take_mono_left:
    31.5       "[| i le j; xs \<in> list(A); j \<in> nat |]
    31.6        ==> <take(i, xs), take(j, xs)> \<in> prefix(A)"
    31.7 -by (blast intro: Nat.le_in_nat take_mono_left_lemma) 
    31.8 +by (blast intro: Nat_ZF.le_in_nat take_mono_left_lemma) 
    31.9  
   31.10  lemma take_mono_right:
   31.11       "[| <xs,ys> \<in> prefix(A); i \<in> nat |] 
    32.1 --- a/src/ZF/UNITY/State.thy	Mon Feb 11 15:19:17 2008 +0100
    32.2 +++ b/src/ZF/UNITY/State.thy	Mon Feb 11 15:40:21 2008 +0100
    32.3 @@ -11,7 +11,7 @@
    32.4  
    32.5  header{*UNITY Program States*}
    32.6  
    32.7 -theory State imports Main begin
    32.8 +theory State imports Main_ZF begin
    32.9  
   32.10  consts var :: i
   32.11  datatype var = Var("i \<in> list(nat)")
    33.1 --- a/src/ZF/Zorn.thy	Mon Feb 11 15:19:17 2008 +0100
    33.2 +++ b/src/ZF/Zorn.thy	Mon Feb 11 15:40:21 2008 +0100
    33.3 @@ -7,7 +7,7 @@
    33.4  
    33.5  header{*Zorn's Lemma*}
    33.6  
    33.7 -theory Zorn imports OrderArith AC Inductive begin
    33.8 +theory Zorn imports OrderArith AC Inductive_ZF begin
    33.9  
   33.10  text{*Based upon the unpublished article ``Towards the Mechanization of the
   33.11  Proofs of Some Classical Theorems of Set Theory,'' by Abrial and Laffitte.*}
    34.1 --- a/src/ZF/ind_syntax.ML	Mon Feb 11 15:19:17 2008 +0100
    34.2 +++ b/src/ZF/ind_syntax.ML	Mon Feb 11 15:40:21 2008 +0100
    34.3 @@ -127,7 +127,7 @@
    34.4            (fn t => "Datatype set not previously declared as constant: " ^
    34.5                     Sign.string_of_term @{theory IFOL} t);
    34.6          val rec_names = (*nat doesn't have to be added*)
    34.7 -	    "Nat.nat" :: map (#1 o dest_Const) rec_hds
    34.8 +	    @{const_name "Nat_ZF.nat"} :: map (#1 o dest_Const) rec_hds
    34.9  	val u = if co then quniv else univ
   34.10  	val cs = (fold o fold) (fn (_, _, _, prems) => prems |> (fold o fold_aterms)
   34.11            (fn t as Const (a, _) => if a mem_string rec_names then I else insert (op =) t
    35.1 --- a/src/ZF/int_arith.ML	Mon Feb 11 15:19:17 2008 +0100
    35.2 +++ b/src/ZF/int_arith.ML	Mon Feb 11 15:40:21 2008 +0100
    35.3 @@ -95,12 +95,12 @@
    35.4  
    35.5  (*Utilities*)
    35.6  
    35.7 -val integ_of_const = Const ("Bin.integ_of", iT --> iT);
    35.8 +val integ_of_const = Const (@{const_name "Bin.integ_of"}, iT --> iT);
    35.9  
   35.10  fun mk_numeral n = integ_of_const $ NumeralSyntax.mk_bin n;
   35.11  
   35.12  (*Decodes a binary INTEGER*)
   35.13 -fun dest_numeral (Const("Bin.integ_of", _) $ w) =
   35.14 +fun dest_numeral (Const(@{const_name "Bin.integ_of"}, _) $ w) =
   35.15       (NumeralSyntax.dest_bin w
   35.16        handle Match => raise TERM("Int_Numeral_Simprocs.dest_numeral:1", [w]))
   35.17    | dest_numeral t =  raise TERM("Int_Numeral_Simprocs.dest_numeral:2", [t]);
   35.18 @@ -111,11 +111,11 @@
   35.19    | find_first_numeral past [] = raise TERM("find_first_numeral", []);
   35.20  
   35.21  val zero = mk_numeral 0;
   35.22 -val mk_plus = FOLogic.mk_binop "Int.zadd";
   35.23 +val mk_plus = FOLogic.mk_binop @{const_name "Int_ZF.zadd"};
   35.24  
   35.25  val iT = Ind_Syntax.iT;
   35.26  
   35.27 -val zminus_const = Const ("Int.zminus", iT --> iT);
   35.28 +val zminus_const = Const (@{const_name "Int_ZF.zminus"}, iT --> iT);
   35.29  
   35.30  (*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*)
   35.31  fun mk_sum []        = zero
   35.32 @@ -126,30 +126,30 @@
   35.33  fun long_mk_sum []        = zero
   35.34    | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
   35.35  
   35.36 -val dest_plus = FOLogic.dest_bin "Int.zadd" iT;
   35.37 +val dest_plus = FOLogic.dest_bin @{const_name "Int_ZF.zadd"} iT;
   35.38  
   35.39  (*decompose additions AND subtractions as a sum*)
   35.40 -fun dest_summing (pos, Const ("Int.zadd", _) $ t $ u, ts) =
   35.41 +fun dest_summing (pos, Const (@{const_name "Int_ZF.zadd"}, _) $ t $ u, ts) =
   35.42          dest_summing (pos, t, dest_summing (pos, u, ts))
   35.43 -  | dest_summing (pos, Const ("Int.zdiff", _) $ t $ u, ts) =
   35.44 +  | dest_summing (pos, Const (@{const_name "Int_ZF.zdiff"}, _) $ t $ u, ts) =
   35.45          dest_summing (pos, t, dest_summing (not pos, u, ts))
   35.46    | dest_summing (pos, t, ts) =
   35.47          if pos then t::ts else zminus_const$t :: ts;
   35.48  
   35.49  fun dest_sum t = dest_summing (true, t, []);
   35.50  
   35.51 -val mk_diff = FOLogic.mk_binop "Int.zdiff";
   35.52 -val dest_diff = FOLogic.dest_bin "Int.zdiff" iT;
   35.53 +val mk_diff = FOLogic.mk_binop @{const_name "Int_ZF.zdiff"};
   35.54 +val dest_diff = FOLogic.dest_bin @{const_name "Int_ZF.zdiff"} iT;
   35.55  
   35.56  val one = mk_numeral 1;
   35.57 -val mk_times = FOLogic.mk_binop "Int.zmult";
   35.58 +val mk_times = FOLogic.mk_binop @{const_name "Int_ZF.zmult"};
   35.59  
   35.60  fun mk_prod [] = one
   35.61    | mk_prod [t] = t
   35.62    | mk_prod (t :: ts) = if t = one then mk_prod ts
   35.63                          else mk_times (t, mk_prod ts);
   35.64  
   35.65 -val dest_times = FOLogic.dest_bin "Int.zmult" iT;
   35.66 +val dest_times = FOLogic.dest_bin @{const_name "Int_ZF.zmult"} iT;
   35.67  
   35.68  fun dest_prod t =
   35.69        let val (t,u) = dest_times t
   35.70 @@ -160,7 +160,7 @@
   35.71  fun mk_coeff (k, t) = mk_times (mk_numeral k, t);
   35.72  
   35.73  (*Express t as a product of (possibly) a numeral with other sorted terms*)
   35.74 -fun dest_coeff sign (Const ("Int.zminus", _) $ t) = dest_coeff (~sign) t
   35.75 +fun dest_coeff sign (Const (@{const_name "Int_ZF.zminus"}, _) $ t) = dest_coeff (~sign) t
   35.76    | dest_coeff sign t =
   35.77      let val ts = sort Term.term_ord (dest_prod t)
   35.78          val (n, ts') = find_first_numeral [] ts
   35.79 @@ -254,8 +254,8 @@
   35.80  structure LessCancelNumerals = CancelNumeralsFun
   35.81   (open CancelNumeralsCommon
   35.82    val prove_conv = ArithData.prove_conv "intless_cancel_numerals"
   35.83 -  val mk_bal   = FOLogic.mk_binrel "Int.zless"
   35.84 -  val dest_bal = FOLogic.dest_bin "Int.zless" iT
   35.85 +  val mk_bal   = FOLogic.mk_binrel @{const_name "Int_ZF.zless"}
   35.86 +  val dest_bal = FOLogic.dest_bin @{const_name "Int_ZF.zless"} iT
   35.87    val bal_add1 = less_add_iff1 RS iff_trans
   35.88    val bal_add2 = less_add_iff2 RS iff_trans
   35.89  );
   35.90 @@ -263,8 +263,8 @@
   35.91  structure LeCancelNumerals = CancelNumeralsFun
   35.92   (open CancelNumeralsCommon
   35.93    val prove_conv = ArithData.prove_conv "intle_cancel_numerals"
   35.94 -  val mk_bal   = FOLogic.mk_binrel "Int.zle"
   35.95 -  val dest_bal = FOLogic.dest_bin "Int.zle" iT
   35.96 +  val mk_bal   = FOLogic.mk_binrel @{const_name "Int_ZF.zle"}
   35.97 +  val dest_bal = FOLogic.dest_bin @{const_name "Int_ZF.zle"} iT
   35.98    val bal_add1 = le_add_iff1 RS iff_trans
   35.99    val bal_add2 = le_add_iff2 RS iff_trans
  35.100  );