diff --git a/src/HOL/Code_Numeral.thy b/src/HOL/Code_Numeral.thy --- a/src/HOL/Code_Numeral.thy +++ b/src/HOL/Code_Numeral.thy @@ -1,1143 +1,1150 @@ (* Title: HOL/Code_Numeral.thy Author: Florian Haftmann, TU Muenchen *) section \Numeric types for code generation onto target language numerals only\ theory Code_Numeral imports Divides Lifting begin subsection \Type of target language integers\ typedef integer = "UNIV :: int set" morphisms int_of_integer integer_of_int .. setup_lifting type_definition_integer lemma integer_eq_iff: "k = l \ int_of_integer k = int_of_integer l" by transfer rule lemma integer_eqI: "int_of_integer k = int_of_integer l \ k = l" using integer_eq_iff [of k l] by simp lemma int_of_integer_integer_of_int [simp]: "int_of_integer (integer_of_int k) = k" by transfer rule lemma integer_of_int_int_of_integer [simp]: "integer_of_int (int_of_integer k) = k" by transfer rule instantiation integer :: ring_1 begin lift_definition zero_integer :: integer is "0 :: int" . declare zero_integer.rep_eq [simp] lift_definition one_integer :: integer is "1 :: int" . declare one_integer.rep_eq [simp] lift_definition plus_integer :: "integer \ integer \ integer" is "plus :: int \ int \ int" . declare plus_integer.rep_eq [simp] lift_definition uminus_integer :: "integer \ integer" is "uminus :: int \ int" . declare uminus_integer.rep_eq [simp] lift_definition minus_integer :: "integer \ integer \ integer" is "minus :: int \ int \ int" . declare minus_integer.rep_eq [simp] lift_definition times_integer :: "integer \ integer \ integer" is "times :: int \ int \ int" . declare times_integer.rep_eq [simp] instance proof qed (transfer, simp add: algebra_simps)+ end instance integer :: Rings.dvd .. -lemma [transfer_rule]: - "rel_fun pcr_integer (rel_fun pcr_integer HOL.iff) Rings.dvd Rings.dvd" - unfolding dvd_def by transfer_prover +context + includes lifting_syntax + notes transfer_rule_numeral [transfer_rule] +begin lemma [transfer_rule]: - "rel_fun (=) pcr_integer (of_bool :: bool \ int) (of_bool :: bool \ integer)" + "(pcr_integer ===> pcr_integer ===> (\)) (dvd) (dvd)" + by (unfold dvd_def) transfer_prover + +lemma [transfer_rule]: + "((\) ===> pcr_integer) of_bool of_bool" by (unfold of_bool_def [abs_def]) transfer_prover lemma [transfer_rule]: - "rel_fun (=) pcr_integer (of_nat :: nat \ int) (of_nat :: nat \ integer)" + "((=) ===> pcr_integer) int of_nat" by (rule transfer_rule_of_nat) transfer_prover+ lemma [transfer_rule]: - "rel_fun (=) pcr_integer (\k :: int. k :: int) (of_int :: int \ integer)" + "((=) ===> pcr_integer) (\k. k) of_int" proof - - have "rel_fun HOL.eq pcr_integer (of_int :: int \ int) (of_int :: int \ integer)" + have "((=) ===> pcr_integer) of_int of_int" by (rule transfer_rule_of_int) transfer_prover+ then show ?thesis by (simp add: id_def) qed lemma [transfer_rule]: - "rel_fun HOL.eq pcr_integer (numeral :: num \ int) (numeral :: num \ integer)" - by (rule transfer_rule_numeral) transfer_prover+ + "((=) ===> pcr_integer) numeral numeral" + by transfer_prover lemma [transfer_rule]: - "rel_fun HOL.eq (rel_fun HOL.eq pcr_integer) (Num.sub :: _ \ _ \ int) (Num.sub :: _ \ _ \ integer)" + "((=) ===> (=) ===> pcr_integer) Num.sub Num.sub" by (unfold Num.sub_def [abs_def]) transfer_prover lemma [transfer_rule]: - "rel_fun pcr_integer (rel_fun (=) pcr_integer) (power :: _ \ _ \ int) (power :: _ \ _ \ integer)" + "(pcr_integer ===> (=) ===> pcr_integer) (^) (^)" by (unfold power_def [abs_def]) transfer_prover +end + lemma int_of_integer_of_nat [simp]: "int_of_integer (of_nat n) = of_nat n" by transfer rule lift_definition integer_of_nat :: "nat \ integer" is "of_nat :: nat \ int" . lemma integer_of_nat_eq_of_nat [code]: "integer_of_nat = of_nat" by transfer rule lemma int_of_integer_integer_of_nat [simp]: "int_of_integer (integer_of_nat n) = of_nat n" by transfer rule lift_definition nat_of_integer :: "integer \ nat" is Int.nat . lemma nat_of_integer_of_nat [simp]: "nat_of_integer (of_nat n) = n" by transfer simp lemma int_of_integer_of_int [simp]: "int_of_integer (of_int k) = k" by transfer simp lemma nat_of_integer_integer_of_nat [simp]: "nat_of_integer (integer_of_nat n) = n" by transfer simp lemma integer_of_int_eq_of_int [simp, code_abbrev]: "integer_of_int = of_int" by transfer (simp add: fun_eq_iff) lemma of_int_integer_of [simp]: "of_int (int_of_integer k) = (k :: integer)" by transfer rule lemma int_of_integer_numeral [simp]: "int_of_integer (numeral k) = numeral k" by transfer rule lemma int_of_integer_sub [simp]: "int_of_integer (Num.sub k l) = Num.sub k l" by transfer rule definition integer_of_num :: "num \ integer" where [simp]: "integer_of_num = numeral" lemma integer_of_num [code]: "integer_of_num Num.One = 1" "integer_of_num (Num.Bit0 n) = (let k = integer_of_num n in k + k)" "integer_of_num (Num.Bit1 n) = (let k = integer_of_num n in k + k + 1)" by (simp_all only: integer_of_num_def numeral.simps Let_def) lemma integer_of_num_triv: "integer_of_num Num.One = 1" "integer_of_num (Num.Bit0 Num.One) = 2" by simp_all instantiation integer :: "{linordered_idom, equal}" begin lift_definition abs_integer :: "integer \ integer" is "abs :: int \ int" . declare abs_integer.rep_eq [simp] lift_definition sgn_integer :: "integer \ integer" is "sgn :: int \ int" . declare sgn_integer.rep_eq [simp] lift_definition less_eq_integer :: "integer \ integer \ bool" is "less_eq :: int \ int \ bool" . lemma integer_less_eq_iff: "k \ l \ int_of_integer k \ int_of_integer l" by (fact less_eq_integer.rep_eq) lift_definition less_integer :: "integer \ integer \ bool" is "less :: int \ int \ bool" . lemma integer_less_iff: "k < l \ int_of_integer k < int_of_integer l" by (fact less_integer.rep_eq) lift_definition equal_integer :: "integer \ integer \ bool" is "HOL.equal :: int \ int \ bool" . instance by standard (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] mult_strict_right_mono linear)+ end lemma [transfer_rule]: "rel_fun pcr_integer (rel_fun pcr_integer pcr_integer) (min :: _ \ _ \ int) (min :: _ \ _ \ integer)" by (unfold min_def [abs_def]) transfer_prover lemma [transfer_rule]: "rel_fun pcr_integer (rel_fun pcr_integer pcr_integer) (max :: _ \ _ \ int) (max :: _ \ _ \ integer)" by (unfold max_def [abs_def]) transfer_prover lemma int_of_integer_min [simp]: "int_of_integer (min k l) = min (int_of_integer k) (int_of_integer l)" by transfer rule lemma int_of_integer_max [simp]: "int_of_integer (max k l) = max (int_of_integer k) (int_of_integer l)" by transfer rule lemma nat_of_integer_non_positive [simp]: "k \ 0 \ nat_of_integer k = 0" by transfer simp lemma of_nat_of_integer [simp]: "of_nat (nat_of_integer k) = max 0 k" by transfer auto instantiation integer :: unique_euclidean_ring begin lift_definition divide_integer :: "integer \ integer \ integer" is "divide :: int \ int \ int" . declare divide_integer.rep_eq [simp] lift_definition modulo_integer :: "integer \ integer \ integer" is "modulo :: int \ int \ int" . declare modulo_integer.rep_eq [simp] lift_definition euclidean_size_integer :: "integer \ nat" is "euclidean_size :: int \ nat" . declare euclidean_size_integer.rep_eq [simp] lift_definition division_segment_integer :: "integer \ integer" is "division_segment :: int \ int" . declare division_segment_integer.rep_eq [simp] instance by (standard; transfer) (use mult_le_mono2 [of 1] in \auto simp add: sgn_mult_abs abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib division_segment_mult division_segment_mod intro: div_eqI\) end lemma [code]: "euclidean_size = nat_of_integer \ abs" by (simp add: fun_eq_iff nat_of_integer.rep_eq) lemma [code]: "division_segment (k :: integer) = (if k \ 0 then 1 else - 1)" by transfer (simp add: division_segment_int_def) instance integer :: unique_euclidean_ring_with_nat by (standard; transfer) (simp_all add: of_nat_div division_segment_int_def) lemma [transfer_rule]: "rel_fun (=) (rel_fun pcr_integer pcr_integer) (push_bit :: _ \ _ \ int) (push_bit :: _ \ _ \ integer)" by (unfold push_bit_eq_mult [abs_def]) transfer_prover lemma [transfer_rule]: "rel_fun (=) (rel_fun pcr_integer pcr_integer) (take_bit :: _ \ _ \ int) (take_bit :: _ \ _ \ integer)" by (unfold take_bit_eq_mod [abs_def]) transfer_prover lemma [transfer_rule]: "rel_fun (=) (rel_fun pcr_integer pcr_integer) (drop_bit :: _ \ _ \ int) (drop_bit :: _ \ _ \ integer)" by (unfold drop_bit_eq_div [abs_def]) transfer_prover instantiation integer :: unique_euclidean_semiring_numeral begin definition divmod_integer :: "num \ num \ integer \ integer" where divmod_integer'_def: "divmod_integer m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_integer :: "num \ integer \ integer \ integer \ integer" where "divmod_step_integer l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" instance proof show "divmod m n = (numeral m div numeral n :: integer, numeral m mod numeral n)" for m n by (fact divmod_integer'_def) show "divmod_step l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" for l and qr :: "integer \ integer" by (fact divmod_step_integer_def) qed (transfer, fact le_add_diff_inverse2 unique_euclidean_semiring_numeral_class.div_less unique_euclidean_semiring_numeral_class.mod_less unique_euclidean_semiring_numeral_class.div_positive unique_euclidean_semiring_numeral_class.mod_less_eq_dividend unique_euclidean_semiring_numeral_class.pos_mod_bound unique_euclidean_semiring_numeral_class.pos_mod_sign unique_euclidean_semiring_numeral_class.mod_mult2_eq unique_euclidean_semiring_numeral_class.div_mult2_eq unique_euclidean_semiring_numeral_class.discrete)+ end declare divmod_algorithm_code [where ?'a = integer, folded integer_of_num_def, unfolded integer_of_num_triv, code] lemma integer_of_nat_0: "integer_of_nat 0 = 0" by transfer simp lemma integer_of_nat_1: "integer_of_nat 1 = 1" by transfer simp lemma integer_of_nat_numeral: "integer_of_nat (numeral n) = numeral n" by transfer simp subsection \Code theorems for target language integers\ text \Constructors\ definition Pos :: "num \ integer" where [simp, code_post]: "Pos = numeral" lemma [transfer_rule]: "rel_fun HOL.eq pcr_integer numeral Pos" by simp transfer_prover lemma Pos_fold [code_unfold]: "numeral Num.One = Pos Num.One" "numeral (Num.Bit0 k) = Pos (Num.Bit0 k)" "numeral (Num.Bit1 k) = Pos (Num.Bit1 k)" by simp_all definition Neg :: "num \ integer" where [simp, code_abbrev]: "Neg n = - Pos n" lemma [transfer_rule]: "rel_fun HOL.eq pcr_integer (\n. - numeral n) Neg" by (simp add: Neg_def [abs_def]) transfer_prover code_datatype "0::integer" Pos Neg text \A further pair of constructors for generated computations\ context begin qualified definition positive :: "num \ integer" where [simp]: "positive = numeral" qualified definition negative :: "num \ integer" where [simp]: "negative = uminus \ numeral" lemma [code_computation_unfold]: "numeral = positive" "Pos = positive" "Neg = negative" by (simp_all add: fun_eq_iff) end text \Auxiliary operations\ lift_definition dup :: "integer \ integer" is "\k::int. k + k" . lemma dup_code [code]: "dup 0 = 0" "dup (Pos n) = Pos (Num.Bit0 n)" "dup (Neg n) = Neg (Num.Bit0 n)" by (transfer, simp only: numeral_Bit0 minus_add_distrib)+ lift_definition sub :: "num \ num \ integer" is "\m n. numeral m - numeral n :: int" . lemma sub_code [code]: "sub Num.One Num.One = 0" "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)" "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)" "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)" "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)" "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)" "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)" "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1" "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1" by (transfer, simp add: dbl_def dbl_inc_def dbl_dec_def)+ text \Implementations\ lemma one_integer_code [code, code_unfold]: "1 = Pos Num.One" by simp lemma plus_integer_code [code]: "k + 0 = (k::integer)" "0 + l = (l::integer)" "Pos m + Pos n = Pos (m + n)" "Pos m + Neg n = sub m n" "Neg m + Pos n = sub n m" "Neg m + Neg n = Neg (m + n)" by (transfer, simp)+ lemma uminus_integer_code [code]: "uminus 0 = (0::integer)" "uminus (Pos m) = Neg m" "uminus (Neg m) = Pos m" by simp_all lemma minus_integer_code [code]: "k - 0 = (k::integer)" "0 - l = uminus (l::integer)" "Pos m - Pos n = sub m n" "Pos m - Neg n = Pos (m + n)" "Neg m - Pos n = Neg (m + n)" "Neg m - Neg n = sub n m" by (transfer, simp)+ lemma abs_integer_code [code]: "\k\ = (if (k::integer) < 0 then - k else k)" by simp lemma sgn_integer_code [code]: "sgn k = (if k = 0 then 0 else if (k::integer) < 0 then - 1 else 1)" by simp lemma times_integer_code [code]: "k * 0 = (0::integer)" "0 * l = (0::integer)" "Pos m * Pos n = Pos (m * n)" "Pos m * Neg n = Neg (m * n)" "Neg m * Pos n = Neg (m * n)" "Neg m * Neg n = Pos (m * n)" by simp_all definition divmod_integer :: "integer \ integer \ integer \ integer" where "divmod_integer k l = (k div l, k mod l)" lemma fst_divmod_integer [simp]: "fst (divmod_integer k l) = k div l" by (simp add: divmod_integer_def) lemma snd_divmod_integer [simp]: "snd (divmod_integer k l) = k mod l" by (simp add: divmod_integer_def) definition divmod_abs :: "integer \ integer \ integer \ integer" where "divmod_abs k l = (\k\ div \l\, \k\ mod \l\)" lemma fst_divmod_abs [simp]: "fst (divmod_abs k l) = \k\ div \l\" by (simp add: divmod_abs_def) lemma snd_divmod_abs [simp]: "snd (divmod_abs k l) = \k\ mod \l\" by (simp add: divmod_abs_def) lemma divmod_abs_code [code]: "divmod_abs (Pos k) (Pos l) = divmod k l" "divmod_abs (Neg k) (Neg l) = divmod k l" "divmod_abs (Neg k) (Pos l) = divmod k l" "divmod_abs (Pos k) (Neg l) = divmod k l" "divmod_abs j 0 = (0, \j\)" "divmod_abs 0 j = (0, 0)" by (simp_all add: prod_eq_iff) lemma divmod_integer_eq_cases: "divmod_integer k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else (apsnd \ times \ sgn) l (if sgn k = sgn l then divmod_abs k l else (let (r, s) = divmod_abs k l in if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" proof - have *: "sgn k = sgn l \ k = 0 \ l = 0 \ 0 < l \ 0 < k \ l < 0 \ k < 0" for k l :: int by (auto simp add: sgn_if) have **: "- k = l * q \ k = - (l * q)" for k l q :: int by auto show ?thesis by (simp add: divmod_integer_def divmod_abs_def) (transfer, auto simp add: * ** not_less zdiv_zminus1_eq_if zmod_zminus1_eq_if div_minus_right mod_minus_right) qed lemma divmod_integer_code [code]: \<^marker>\contributor \René Thiemann\\ \<^marker>\contributor \Akihisa Yamada\\ "divmod_integer k l = (if k = 0 then (0, 0) else if l > 0 then (if k > 0 then Code_Numeral.divmod_abs k l else case Code_Numeral.divmod_abs k l of (r, s) \ if s = 0 then (- r, 0) else (- r - 1, l - s)) else if l = 0 then (0, k) else apsnd uminus (if k < 0 then Code_Numeral.divmod_abs k l else case Code_Numeral.divmod_abs k l of (r, s) \ if s = 0 then (- r, 0) else (- r - 1, - l - s)))" by (cases l "0 :: integer" rule: linorder_cases) (auto split: prod.splits simp add: divmod_integer_eq_cases) lemma div_integer_code [code]: "k div l = fst (divmod_integer k l)" by simp lemma mod_integer_code [code]: "k mod l = snd (divmod_integer k l)" by simp definition bit_cut_integer :: "integer \ integer \ bool" where "bit_cut_integer k = (k div 2, odd k)" lemma bit_cut_integer_code [code]: "bit_cut_integer k = (if k = 0 then (0, False) else let (r, s) = Code_Numeral.divmod_abs k 2 in (if k > 0 then r else - r - s, s = 1))" proof - have "bit_cut_integer k = (let (r, s) = divmod_integer k 2 in (r, s = 1))" by (simp add: divmod_integer_def bit_cut_integer_def odd_iff_mod_2_eq_one) then show ?thesis by (simp add: divmod_integer_code) (auto simp add: split_def) qed lemma equal_integer_code [code]: "HOL.equal 0 (0::integer) \ True" "HOL.equal 0 (Pos l) \ False" "HOL.equal 0 (Neg l) \ False" "HOL.equal (Pos k) 0 \ False" "HOL.equal (Pos k) (Pos l) \ HOL.equal k l" "HOL.equal (Pos k) (Neg l) \ False" "HOL.equal (Neg k) 0 \ False" "HOL.equal (Neg k) (Pos l) \ False" "HOL.equal (Neg k) (Neg l) \ HOL.equal k l" by (simp_all add: equal) lemma equal_integer_refl [code nbe]: "HOL.equal (k::integer) k \ True" by (fact equal_refl) lemma less_eq_integer_code [code]: "0 \ (0::integer) \ True" "0 \ Pos l \ True" "0 \ Neg l \ False" "Pos k \ 0 \ False" "Pos k \ Pos l \ k \ l" "Pos k \ Neg l \ False" "Neg k \ 0 \ True" "Neg k \ Pos l \ True" "Neg k \ Neg l \ l \ k" by simp_all lemma less_integer_code [code]: "0 < (0::integer) \ False" "0 < Pos l \ True" "0 < Neg l \ False" "Pos k < 0 \ False" "Pos k < Pos l \ k < l" "Pos k < Neg l \ False" "Neg k < 0 \ True" "Neg k < Pos l \ True" "Neg k < Neg l \ l < k" by simp_all lift_definition num_of_integer :: "integer \ num" is "num_of_nat \ nat" . lemma num_of_integer_code [code]: "num_of_integer k = (if k \ 1 then Num.One else let (l, j) = divmod_integer k 2; l' = num_of_integer l; l'' = l' + l' in if j = 0 then l'' else l'' + Num.One)" proof - { assume "int_of_integer k mod 2 = 1" then have "nat (int_of_integer k mod 2) = nat 1" by simp moreover assume *: "1 < int_of_integer k" ultimately have **: "nat (int_of_integer k) mod 2 = 1" by (simp add: nat_mod_distrib) have "num_of_nat (nat (int_of_integer k)) = num_of_nat (2 * (nat (int_of_integer k) div 2) + nat (int_of_integer k) mod 2)" by simp then have "num_of_nat (nat (int_of_integer k)) = num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + nat (int_of_integer k) mod 2)" by (simp add: mult_2) with ** have "num_of_nat (nat (int_of_integer k)) = num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + 1)" by simp } note aux = this show ?thesis by (auto simp add: num_of_integer_def nat_of_integer_def Let_def case_prod_beta not_le integer_eq_iff less_eq_integer_def nat_mult_distrib nat_div_distrib num_of_nat_One num_of_nat_plus_distrib mult_2 [where 'a=nat] aux add_One) qed lemma nat_of_integer_code [code]: "nat_of_integer k = (if k \ 0 then 0 else let (l, j) = divmod_integer k 2; l' = nat_of_integer l; l'' = l' + l' in if j = 0 then l'' else l'' + 1)" proof - obtain j where k: "k = integer_of_int j" proof show "k = integer_of_int (int_of_integer k)" by simp qed have *: "nat j mod 2 = nat_of_integer (of_int j mod 2)" if "j \ 0" using that by transfer (simp add: nat_mod_distrib) from k show ?thesis by (auto simp add: split_def Let_def nat_of_integer_def nat_div_distrib mult_2 [symmetric] minus_mod_eq_mult_div [symmetric] *) qed lemma int_of_integer_code [code]: "int_of_integer k = (if k < 0 then - (int_of_integer (- k)) else if k = 0 then 0 else let (l, j) = divmod_integer k 2; l' = 2 * int_of_integer l in if j = 0 then l' else l' + 1)" by (auto simp add: split_def Let_def integer_eq_iff minus_mod_eq_mult_div [symmetric]) lemma integer_of_int_code [code]: "integer_of_int k = (if k < 0 then - (integer_of_int (- k)) else if k = 0 then 0 else let l = 2 * integer_of_int (k div 2); j = k mod 2 in if j = 0 then l else l + 1)" by (auto simp add: split_def Let_def integer_eq_iff minus_mod_eq_mult_div [symmetric]) hide_const (open) Pos Neg sub dup divmod_abs subsection \Serializer setup for target language integers\ code_reserved Eval int Integer abs code_printing type_constructor integer \ (SML) "IntInf.int" and (OCaml) "Z.t" and (Haskell) "Integer" and (Scala) "BigInt" and (Eval) "int" | class_instance integer :: equal \ (Haskell) - code_printing constant "0::integer" \ (SML) "!(0/ :/ IntInf.int)" and (OCaml) "Z.zero" and (Haskell) "!(0/ ::/ Integer)" and (Scala) "BigInt(0)" setup \ fold (fn target => Numeral.add_code \<^const_name>\Code_Numeral.Pos\ I Code_Printer.literal_numeral target #> Numeral.add_code \<^const_name>\Code_Numeral.Neg\ (~) Code_Printer.literal_numeral target) ["SML", "OCaml", "Haskell", "Scala"] \ code_printing constant "plus :: integer \ _ \ _" \ (SML) "IntInf.+ ((_), (_))" and (OCaml) "Z.add" and (Haskell) infixl 6 "+" and (Scala) infixl 7 "+" and (Eval) infixl 8 "+" | constant "uminus :: integer \ _" \ (SML) "IntInf.~" and (OCaml) "Z.neg" and (Haskell) "negate" and (Scala) "!(- _)" and (Eval) "~/ _" | constant "minus :: integer \ _" \ (SML) "IntInf.- ((_), (_))" and (OCaml) "Z.sub" and (Haskell) infixl 6 "-" and (Scala) infixl 7 "-" and (Eval) infixl 8 "-" | constant Code_Numeral.dup \ (SML) "IntInf.*/ (2,/ (_))" and (OCaml) "Z.shift'_left/ _/ 1" and (Haskell) "!(2 * _)" and (Scala) "!(2 * _)" and (Eval) "!(2 * _)" | constant Code_Numeral.sub \ (SML) "!(raise/ Fail/ \"sub\")" and (OCaml) "failwith/ \"sub\"" and (Haskell) "error/ \"sub\"" and (Scala) "!sys.error(\"sub\")" | constant "times :: integer \ _ \ _" \ (SML) "IntInf.* ((_), (_))" and (OCaml) "Z.mul" and (Haskell) infixl 7 "*" and (Scala) infixl 8 "*" and (Eval) infixl 9 "*" | constant Code_Numeral.divmod_abs \ (SML) "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)" and (OCaml) "!(fun k l ->/ if Z.equal Z.zero l then/ (Z.zero, l) else/ Z.div'_rem/ (Z.abs k)/ (Z.abs l))" and (Haskell) "divMod/ (abs _)/ (abs _)" and (Scala) "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))" and (Eval) "Integer.div'_mod/ (abs _)/ (abs _)" | constant "HOL.equal :: integer \ _ \ bool" \ (SML) "!((_ : IntInf.int) = _)" and (OCaml) "Z.equal" and (Haskell) infix 4 "==" and (Scala) infixl 5 "==" and (Eval) infixl 6 "=" | constant "less_eq :: integer \ _ \ bool" \ (SML) "IntInf.<= ((_), (_))" and (OCaml) "Z.leq" and (Haskell) infix 4 "<=" and (Scala) infixl 4 "<=" and (Eval) infixl 6 "<=" | constant "less :: integer \ _ \ bool" \ (SML) "IntInf.< ((_), (_))" and (OCaml) "Z.lt" and (Haskell) infix 4 "<" and (Scala) infixl 4 "<" and (Eval) infixl 6 "<" | constant "abs :: integer \ _" \ (SML) "IntInf.abs" and (OCaml) "Z.abs" and (Haskell) "Prelude.abs" and (Scala) "_.abs" and (Eval) "abs" code_identifier code_module Code_Numeral \ (SML) Arith and (OCaml) Arith and (Haskell) Arith subsection \Type of target language naturals\ typedef natural = "UNIV :: nat set" morphisms nat_of_natural natural_of_nat .. setup_lifting type_definition_natural lemma natural_eq_iff [termination_simp]: "m = n \ nat_of_natural m = nat_of_natural n" by transfer rule lemma natural_eqI: "nat_of_natural m = nat_of_natural n \ m = n" using natural_eq_iff [of m n] by simp lemma nat_of_natural_of_nat_inverse [simp]: "nat_of_natural (natural_of_nat n) = n" by transfer rule lemma natural_of_nat_of_natural_inverse [simp]: "natural_of_nat (nat_of_natural n) = n" by transfer rule instantiation natural :: "{comm_monoid_diff, semiring_1}" begin lift_definition zero_natural :: natural is "0 :: nat" . declare zero_natural.rep_eq [simp] lift_definition one_natural :: natural is "1 :: nat" . declare one_natural.rep_eq [simp] lift_definition plus_natural :: "natural \ natural \ natural" is "plus :: nat \ nat \ nat" . declare plus_natural.rep_eq [simp] lift_definition minus_natural :: "natural \ natural \ natural" is "minus :: nat \ nat \ nat" . declare minus_natural.rep_eq [simp] lift_definition times_natural :: "natural \ natural \ natural" is "times :: nat \ nat \ nat" . declare times_natural.rep_eq [simp] instance proof qed (transfer, simp add: algebra_simps)+ end instance natural :: Rings.dvd .. lemma [transfer_rule]: "rel_fun pcr_natural (rel_fun pcr_natural HOL.iff) Rings.dvd Rings.dvd" unfolding dvd_def by transfer_prover lemma [transfer_rule]: "rel_fun (=) pcr_natural (of_bool :: bool \ nat) (of_bool :: bool \ natural)" by (unfold of_bool_def [abs_def]) transfer_prover lemma [transfer_rule]: "rel_fun HOL.eq pcr_natural (\n::nat. n) (of_nat :: nat \ natural)" proof - have "rel_fun HOL.eq pcr_natural (of_nat :: nat \ nat) (of_nat :: nat \ natural)" by (unfold of_nat_def [abs_def]) transfer_prover then show ?thesis by (simp add: id_def) qed lemma [transfer_rule]: "rel_fun HOL.eq pcr_natural (numeral :: num \ nat) (numeral :: num \ natural)" proof - have "rel_fun HOL.eq pcr_natural (numeral :: num \ nat) (\n. of_nat (numeral n))" by transfer_prover then show ?thesis by simp qed lemma [transfer_rule]: "rel_fun pcr_natural (rel_fun (=) pcr_natural) (power :: _ \ _ \ nat) (power :: _ \ _ \ natural)" by (unfold power_def [abs_def]) transfer_prover lemma nat_of_natural_of_nat [simp]: "nat_of_natural (of_nat n) = n" by transfer rule lemma natural_of_nat_of_nat [simp, code_abbrev]: "natural_of_nat = of_nat" by transfer rule lemma of_nat_of_natural [simp]: "of_nat (nat_of_natural n) = n" by transfer rule lemma nat_of_natural_numeral [simp]: "nat_of_natural (numeral k) = numeral k" by transfer rule instantiation natural :: "{linordered_semiring, equal}" begin lift_definition less_eq_natural :: "natural \ natural \ bool" is "less_eq :: nat \ nat \ bool" . declare less_eq_natural.rep_eq [termination_simp] lift_definition less_natural :: "natural \ natural \ bool" is "less :: nat \ nat \ bool" . declare less_natural.rep_eq [termination_simp] lift_definition equal_natural :: "natural \ natural \ bool" is "HOL.equal :: nat \ nat \ bool" . instance proof qed (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] linear)+ end lemma [transfer_rule]: "rel_fun pcr_natural (rel_fun pcr_natural pcr_natural) (min :: _ \ _ \ nat) (min :: _ \ _ \ natural)" by (unfold min_def [abs_def]) transfer_prover lemma [transfer_rule]: "rel_fun pcr_natural (rel_fun pcr_natural pcr_natural) (max :: _ \ _ \ nat) (max :: _ \ _ \ natural)" by (unfold max_def [abs_def]) transfer_prover lemma nat_of_natural_min [simp]: "nat_of_natural (min k l) = min (nat_of_natural k) (nat_of_natural l)" by transfer rule lemma nat_of_natural_max [simp]: "nat_of_natural (max k l) = max (nat_of_natural k) (nat_of_natural l)" by transfer rule instantiation natural :: unique_euclidean_semiring begin lift_definition divide_natural :: "natural \ natural \ natural" is "divide :: nat \ nat \ nat" . declare divide_natural.rep_eq [simp] lift_definition modulo_natural :: "natural \ natural \ natural" is "modulo :: nat \ nat \ nat" . declare modulo_natural.rep_eq [simp] lift_definition euclidean_size_natural :: "natural \ nat" is "euclidean_size :: nat \ nat" . declare euclidean_size_natural.rep_eq [simp] lift_definition division_segment_natural :: "natural \ natural" is "division_segment :: nat \ nat" . declare division_segment_natural.rep_eq [simp] instance by (standard; transfer) (auto simp add: algebra_simps unit_factor_nat_def gr0_conv_Suc) end lemma [code]: "euclidean_size = nat_of_natural" by (simp add: fun_eq_iff) lemma [code]: "division_segment (n::natural) = 1" by (simp add: natural_eq_iff) instance natural :: linordered_semidom by (standard; transfer) simp_all instance natural :: unique_euclidean_semiring_with_nat by (standard; transfer) simp_all lemma [transfer_rule]: "rel_fun (=) (rel_fun pcr_natural pcr_natural) (push_bit :: _ \ _ \ nat) (push_bit :: _ \ _ \ natural)" by (unfold push_bit_eq_mult [abs_def]) transfer_prover lemma [transfer_rule]: "rel_fun (=) (rel_fun pcr_natural pcr_natural) (take_bit :: _ \ _ \ nat) (take_bit :: _ \ _ \ natural)" by (unfold take_bit_eq_mod [abs_def]) transfer_prover lemma [transfer_rule]: "rel_fun (=) (rel_fun pcr_natural pcr_natural) (drop_bit :: _ \ _ \ nat) (drop_bit :: _ \ _ \ natural)" by (unfold drop_bit_eq_div [abs_def]) transfer_prover lift_definition natural_of_integer :: "integer \ natural" is "nat :: int \ nat" . lift_definition integer_of_natural :: "natural \ integer" is "of_nat :: nat \ int" . lemma natural_of_integer_of_natural [simp]: "natural_of_integer (integer_of_natural n) = n" by transfer simp lemma integer_of_natural_of_integer [simp]: "integer_of_natural (natural_of_integer k) = max 0 k" by transfer auto lemma int_of_integer_of_natural [simp]: "int_of_integer (integer_of_natural n) = of_nat (nat_of_natural n)" by transfer rule lemma integer_of_natural_of_nat [simp]: "integer_of_natural (of_nat n) = of_nat n" by transfer rule lemma [measure_function]: "is_measure nat_of_natural" by (rule is_measure_trivial) subsection \Inductive representation of target language naturals\ lift_definition Suc :: "natural \ natural" is Nat.Suc . declare Suc.rep_eq [simp] old_rep_datatype "0::natural" Suc by (transfer, fact nat.induct nat.inject nat.distinct)+ lemma natural_cases [case_names nat, cases type: natural]: fixes m :: natural assumes "\n. m = of_nat n \ P" shows P using assms by transfer blast instantiation natural :: size begin definition size_nat where [simp, code]: "size_nat = nat_of_natural" instance .. end lemma natural_decr [termination_simp]: "n \ 0 \ nat_of_natural n - Nat.Suc 0 < nat_of_natural n" by transfer simp lemma natural_zero_minus_one: "(0::natural) - 1 = 0" by (rule zero_diff) lemma Suc_natural_minus_one: "Suc n - 1 = n" by transfer simp hide_const (open) Suc subsection \Code refinement for target language naturals\ lift_definition Nat :: "integer \ natural" is nat . lemma [code_post]: "Nat 0 = 0" "Nat 1 = 1" "Nat (numeral k) = numeral k" by (transfer, simp)+ lemma [code abstype]: "Nat (integer_of_natural n) = n" by transfer simp lemma [code]: "natural_of_nat n = natural_of_integer (integer_of_nat n)" by transfer simp lemma [code abstract]: "integer_of_natural (natural_of_integer k) = max 0 k" by simp lemma [code_abbrev]: "natural_of_integer (Code_Numeral.Pos k) = numeral k" by transfer simp lemma [code abstract]: "integer_of_natural 0 = 0" by transfer simp lemma [code abstract]: "integer_of_natural 1 = 1" by transfer simp lemma [code abstract]: "integer_of_natural (Code_Numeral.Suc n) = integer_of_natural n + 1" by transfer simp lemma [code]: "nat_of_natural = nat_of_integer \ integer_of_natural" by transfer (simp add: fun_eq_iff) lemma [code, code_unfold]: "case_natural f g n = (if n = 0 then f else g (n - 1))" by (cases n rule: natural.exhaust) (simp_all, simp add: Suc_def) declare natural.rec [code del] lemma [code abstract]: "integer_of_natural (m + n) = integer_of_natural m + integer_of_natural n" by transfer simp lemma [code abstract]: "integer_of_natural (m - n) = max 0 (integer_of_natural m - integer_of_natural n)" by transfer simp lemma [code abstract]: "integer_of_natural (m * n) = integer_of_natural m * integer_of_natural n" by transfer simp lemma [code abstract]: "integer_of_natural (m div n) = integer_of_natural m div integer_of_natural n" by transfer (simp add: zdiv_int) lemma [code abstract]: "integer_of_natural (m mod n) = integer_of_natural m mod integer_of_natural n" by transfer (simp add: zmod_int) lemma [code]: "HOL.equal m n \ HOL.equal (integer_of_natural m) (integer_of_natural n)" by transfer (simp add: equal) lemma [code nbe]: "HOL.equal n (n::natural) \ True" by (rule equal_class.equal_refl) lemma [code]: "m \ n \ integer_of_natural m \ integer_of_natural n" by transfer simp lemma [code]: "m < n \ integer_of_natural m < integer_of_natural n" by transfer simp hide_const (open) Nat lifting_update integer.lifting lifting_forget integer.lifting lifting_update natural.lifting lifting_forget natural.lifting code_reflect Code_Numeral datatypes natural functions "Code_Numeral.Suc" "0 :: natural" "1 :: natural" "plus :: natural \ _" "minus :: natural \ _" "times :: natural \ _" "divide :: natural \ _" "modulo :: natural \ _" integer_of_natural natural_of_integer end diff --git a/src/HOL/Filter.thy b/src/HOL/Filter.thy --- a/src/HOL/Filter.thy +++ b/src/HOL/Filter.thy @@ -1,1999 +1,2006 @@ (* Title: HOL/Filter.thy Author: Brian Huffman Author: Johannes Hölzl *) section \Filters on predicates\ theory Filter imports Set_Interval Lifting_Set begin subsection \Filters\ text \ This definition also allows non-proper filters. \ locale is_filter = fixes F :: "('a \ bool) \ bool" assumes True: "F (\x. True)" assumes conj: "F (\x. P x) \ F (\x. Q x) \ F (\x. P x \ Q x)" assumes mono: "\x. P x \ Q x \ F (\x. P x) \ F (\x. Q x)" typedef 'a filter = "{F :: ('a \ bool) \ bool. is_filter F}" proof show "(\x. True) \ ?filter" by (auto intro: is_filter.intro) qed lemma is_filter_Rep_filter: "is_filter (Rep_filter F)" using Rep_filter [of F] by simp lemma Abs_filter_inverse': assumes "is_filter F" shows "Rep_filter (Abs_filter F) = F" using assms by (simp add: Abs_filter_inverse) subsubsection \Eventually\ definition eventually :: "('a \ bool) \ 'a filter \ bool" where "eventually P F \ Rep_filter F P" syntax "_eventually" :: "pttrn => 'a filter => bool => bool" ("(3\\<^sub>F _ in _./ _)" [0, 0, 10] 10) translations "\\<^sub>Fx in F. P" == "CONST eventually (\x. P) F" lemma eventually_Abs_filter: assumes "is_filter F" shows "eventually P (Abs_filter F) = F P" unfolding eventually_def using assms by (simp add: Abs_filter_inverse) lemma filter_eq_iff: shows "F = F' \ (\P. eventually P F = eventually P F')" unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def .. lemma eventually_True [simp]: "eventually (\x. True) F" unfolding eventually_def by (rule is_filter.True [OF is_filter_Rep_filter]) lemma always_eventually: "\x. P x \ eventually P F" proof - assume "\x. P x" hence "P = (\x. True)" by (simp add: ext) thus "eventually P F" by simp qed lemma eventuallyI: "(\x. P x) \ eventually P F" by (auto intro: always_eventually) lemma eventually_mono: "\eventually P F; \x. P x \ Q x\ \ eventually Q F" unfolding eventually_def by (blast intro: is_filter.mono [OF is_filter_Rep_filter]) lemma eventually_conj: assumes P: "eventually (\x. P x) F" assumes Q: "eventually (\x. Q x) F" shows "eventually (\x. P x \ Q x) F" using assms unfolding eventually_def by (rule is_filter.conj [OF is_filter_Rep_filter]) lemma eventually_mp: assumes "eventually (\x. P x \ Q x) F" assumes "eventually (\x. P x) F" shows "eventually (\x. Q x) F" proof - have "eventually (\x. (P x \ Q x) \ P x) F" using assms by (rule eventually_conj) then show ?thesis by (blast intro: eventually_mono) qed lemma eventually_rev_mp: assumes "eventually (\x. P x) F" assumes "eventually (\x. P x \ Q x) F" shows "eventually (\x. Q x) F" using assms(2) assms(1) by (rule eventually_mp) lemma eventually_conj_iff: "eventually (\x. P x \ Q x) F \ eventually P F \ eventually Q F" by (auto intro: eventually_conj elim: eventually_rev_mp) lemma eventually_elim2: assumes "eventually (\i. P i) F" assumes "eventually (\i. Q i) F" assumes "\i. P i \ Q i \ R i" shows "eventually (\i. R i) F" using assms by (auto elim!: eventually_rev_mp) lemma eventually_ball_finite_distrib: "finite A \ (eventually (\x. \y\A. P x y) net) \ (\y\A. eventually (\x. P x y) net)" by (induction A rule: finite_induct) (auto simp: eventually_conj_iff) lemma eventually_ball_finite: "finite A \ \y\A. eventually (\x. P x y) net \ eventually (\x. \y\A. P x y) net" by (auto simp: eventually_ball_finite_distrib) lemma eventually_all_finite: fixes P :: "'a \ 'b::finite \ bool" assumes "\y. eventually (\x. P x y) net" shows "eventually (\x. \y. P x y) net" using eventually_ball_finite [of UNIV P] assms by simp lemma eventually_ex: "(\\<^sub>Fx in F. \y. P x y) \ (\Y. \\<^sub>Fx in F. P x (Y x))" proof assume "\\<^sub>Fx in F. \y. P x y" then have "\\<^sub>Fx in F. P x (SOME y. P x y)" by (auto intro: someI_ex eventually_mono) then show "\Y. \\<^sub>Fx in F. P x (Y x)" by auto qed (auto intro: eventually_mono) lemma not_eventually_impI: "eventually P F \ \ eventually Q F \ \ eventually (\x. P x \ Q x) F" by (auto intro: eventually_mp) lemma not_eventuallyD: "\ eventually P F \ \x. \ P x" by (metis always_eventually) lemma eventually_subst: assumes "eventually (\n. P n = Q n) F" shows "eventually P F = eventually Q F" (is "?L = ?R") proof - from assms have "eventually (\x. P x \ Q x) F" and "eventually (\x. Q x \ P x) F" by (auto elim: eventually_mono) then show ?thesis by (auto elim: eventually_elim2) qed subsection \ Frequently as dual to eventually \ definition frequently :: "('a \ bool) \ 'a filter \ bool" where "frequently P F \ \ eventually (\x. \ P x) F" syntax "_frequently" :: "pttrn \ 'a filter \ bool \ bool" ("(3\\<^sub>F _ in _./ _)" [0, 0, 10] 10) translations "\\<^sub>Fx in F. P" == "CONST frequently (\x. P) F" lemma not_frequently_False [simp]: "\ (\\<^sub>Fx in F. False)" by (simp add: frequently_def) lemma frequently_ex: "\\<^sub>Fx in F. P x \ \x. P x" by (auto simp: frequently_def dest: not_eventuallyD) lemma frequentlyE: assumes "frequently P F" obtains x where "P x" using frequently_ex[OF assms] by auto lemma frequently_mp: assumes ev: "\\<^sub>Fx in F. P x \ Q x" and P: "\\<^sub>Fx in F. P x" shows "\\<^sub>Fx in F. Q x" proof - from ev have "eventually (\x. \ Q x \ \ P x) F" by (rule eventually_rev_mp) (auto intro!: always_eventually) from eventually_mp[OF this] P show ?thesis by (auto simp: frequently_def) qed lemma frequently_rev_mp: assumes "\\<^sub>Fx in F. P x" assumes "\\<^sub>Fx in F. P x \ Q x" shows "\\<^sub>Fx in F. Q x" using assms(2) assms(1) by (rule frequently_mp) lemma frequently_mono: "(\x. P x \ Q x) \ frequently P F \ frequently Q F" using frequently_mp[of P Q] by (simp add: always_eventually) lemma frequently_elim1: "\\<^sub>Fx in F. P x \ (\i. P i \ Q i) \ \\<^sub>Fx in F. Q x" by (metis frequently_mono) lemma frequently_disj_iff: "(\\<^sub>Fx in F. P x \ Q x) \ (\\<^sub>Fx in F. P x) \ (\\<^sub>Fx in F. Q x)" by (simp add: frequently_def eventually_conj_iff) lemma frequently_disj: "\\<^sub>Fx in F. P x \ \\<^sub>Fx in F. Q x \ \\<^sub>Fx in F. P x \ Q x" by (simp add: frequently_disj_iff) lemma frequently_bex_finite_distrib: assumes "finite A" shows "(\\<^sub>Fx in F. \y\A. P x y) \ (\y\A. \\<^sub>Fx in F. P x y)" using assms by induction (auto simp: frequently_disj_iff) lemma frequently_bex_finite: "finite A \ \\<^sub>Fx in F. \y\A. P x y \ \y\A. \\<^sub>Fx in F. P x y" by (simp add: frequently_bex_finite_distrib) lemma frequently_all: "(\\<^sub>Fx in F. \y. P x y) \ (\Y. \\<^sub>Fx in F. P x (Y x))" using eventually_ex[of "\x y. \ P x y" F] by (simp add: frequently_def) lemma shows not_eventually: "\ eventually P F \ (\\<^sub>Fx in F. \ P x)" and not_frequently: "\ frequently P F \ (\\<^sub>Fx in F. \ P x)" by (auto simp: frequently_def) lemma frequently_imp_iff: "(\\<^sub>Fx in F. P x \ Q x) \ (eventually P F \ frequently Q F)" unfolding imp_conv_disj frequently_disj_iff not_eventually[symmetric] .. lemma eventually_frequently_const_simps: "(\\<^sub>Fx in F. P x \ C) \ (\\<^sub>Fx in F. P x) \ C" "(\\<^sub>Fx in F. C \ P x) \ C \ (\\<^sub>Fx in F. P x)" "(\\<^sub>Fx in F. P x \ C) \ (\\<^sub>Fx in F. P x) \ C" "(\\<^sub>Fx in F. C \ P x) \ C \ (\\<^sub>Fx in F. P x)" "(\\<^sub>Fx in F. P x \ C) \ ((\\<^sub>Fx in F. P x) \ C)" "(\\<^sub>Fx in F. C \ P x) \ (C \ (\\<^sub>Fx in F. P x))" by (cases C; simp add: not_frequently)+ lemmas eventually_frequently_simps = eventually_frequently_const_simps not_eventually eventually_conj_iff eventually_ball_finite_distrib eventually_ex not_frequently frequently_disj_iff frequently_bex_finite_distrib frequently_all frequently_imp_iff ML \ fun eventually_elim_tac facts = CONTEXT_SUBGOAL (fn (goal, i) => fn (ctxt, st) => let val mp_facts = facts RL @{thms eventually_rev_mp} val rule = @{thm eventuallyI} |> fold (fn mp_fact => fn th => th RS mp_fact) mp_facts |> funpow (length facts) (fn th => @{thm impI} RS th) val cases_prop = Thm.prop_of (Rule_Cases.internalize_params (rule RS Goal.init (Thm.cterm_of ctxt goal))) val cases = Rule_Cases.make_common ctxt cases_prop [(("elim", []), [])] in CONTEXT_CASES cases (resolve_tac ctxt [rule] i) (ctxt, st) end) \ method_setup eventually_elim = \ Scan.succeed (fn _ => CONTEXT_METHOD (fn facts => eventually_elim_tac facts 1)) \ "elimination of eventually quantifiers" subsubsection \Finer-than relation\ text \\<^term>\F \ F'\ means that filter \<^term>\F\ is finer than filter \<^term>\F'\.\ instantiation filter :: (type) complete_lattice begin definition le_filter_def: "F \ F' \ (\P. eventually P F' \ eventually P F)" definition "(F :: 'a filter) < F' \ F \ F' \ \ F' \ F" definition "top = Abs_filter (\P. \x. P x)" definition "bot = Abs_filter (\P. True)" definition "sup F F' = Abs_filter (\P. eventually P F \ eventually P F')" definition "inf F F' = Abs_filter (\P. \Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" definition "Sup S = Abs_filter (\P. \F\S. eventually P F)" definition "Inf S = Sup {F::'a filter. \F'\S. F \ F'}" lemma eventually_top [simp]: "eventually P top \ (\x. P x)" unfolding top_filter_def by (rule eventually_Abs_filter, rule is_filter.intro, auto) lemma eventually_bot [simp]: "eventually P bot" unfolding bot_filter_def by (subst eventually_Abs_filter, rule is_filter.intro, auto) lemma eventually_sup: "eventually P (sup F F') \ eventually P F \ eventually P F'" unfolding sup_filter_def by (rule eventually_Abs_filter, rule is_filter.intro) (auto elim!: eventually_rev_mp) lemma eventually_inf: "eventually P (inf F F') \ (\Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" unfolding inf_filter_def apply (rule eventually_Abs_filter, rule is_filter.intro) apply (fast intro: eventually_True) apply clarify apply (intro exI conjI) apply (erule (1) eventually_conj) apply (erule (1) eventually_conj) apply simp apply auto done lemma eventually_Sup: "eventually P (Sup S) \ (\F\S. eventually P F)" unfolding Sup_filter_def apply (rule eventually_Abs_filter, rule is_filter.intro) apply (auto intro: eventually_conj elim!: eventually_rev_mp) done instance proof fix F F' F'' :: "'a filter" and S :: "'a filter set" { show "F < F' \ F \ F' \ \ F' \ F" by (rule less_filter_def) } { show "F \ F" unfolding le_filter_def by simp } { assume "F \ F'" and "F' \ F''" thus "F \ F''" unfolding le_filter_def by simp } { assume "F \ F'" and "F' \ F" thus "F = F'" unfolding le_filter_def filter_eq_iff by fast } { show "inf F F' \ F" and "inf F F' \ F'" unfolding le_filter_def eventually_inf by (auto intro: eventually_True) } { assume "F \ F'" and "F \ F''" thus "F \ inf F' F''" unfolding le_filter_def eventually_inf by (auto intro: eventually_mono [OF eventually_conj]) } { show "F \ sup F F'" and "F' \ sup F F'" unfolding le_filter_def eventually_sup by simp_all } { assume "F \ F''" and "F' \ F''" thus "sup F F' \ F''" unfolding le_filter_def eventually_sup by simp } { assume "F'' \ S" thus "Inf S \ F''" unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } { assume "\F'. F' \ S \ F \ F'" thus "F \ Inf S" unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } { assume "F \ S" thus "F \ Sup S" unfolding le_filter_def eventually_Sup by simp } { assume "\F. F \ S \ F \ F'" thus "Sup S \ F'" unfolding le_filter_def eventually_Sup by simp } { show "Inf {} = (top::'a filter)" by (auto simp: top_filter_def Inf_filter_def Sup_filter_def) (metis (full_types) top_filter_def always_eventually eventually_top) } { show "Sup {} = (bot::'a filter)" by (auto simp: bot_filter_def Sup_filter_def) } qed end instance filter :: (type) distrib_lattice proof fix F G H :: "'a filter" show "sup F (inf G H) = inf (sup F G) (sup F H)" proof (rule order.antisym) show "inf (sup F G) (sup F H) \ sup F (inf G H)" unfolding le_filter_def eventually_sup proof safe fix P assume 1: "eventually P F" and 2: "eventually P (inf G H)" from 2 obtain Q R where QR: "eventually Q G" "eventually R H" "\x. Q x \ R x \ P x" by (auto simp: eventually_inf) define Q' where "Q' = (\x. Q x \ P x)" define R' where "R' = (\x. R x \ P x)" from 1 have "eventually Q' F" by (elim eventually_mono) (auto simp: Q'_def) moreover from 1 have "eventually R' F" by (elim eventually_mono) (auto simp: R'_def) moreover from QR(1) have "eventually Q' G" by (elim eventually_mono) (auto simp: Q'_def) moreover from QR(2) have "eventually R' H" by (elim eventually_mono)(auto simp: R'_def) moreover from QR have "P x" if "Q' x" "R' x" for x using that by (auto simp: Q'_def R'_def) ultimately show "eventually P (inf (sup F G) (sup F H))" by (auto simp: eventually_inf eventually_sup) qed qed (auto intro: inf.coboundedI1 inf.coboundedI2) qed lemma filter_leD: "F \ F' \ eventually P F' \ eventually P F" unfolding le_filter_def by simp lemma filter_leI: "(\P. eventually P F' \ eventually P F) \ F \ F'" unfolding le_filter_def by simp lemma eventually_False: "eventually (\x. False) F \ F = bot" unfolding filter_eq_iff by (auto elim: eventually_rev_mp) lemma eventually_frequently: "F \ bot \ eventually P F \ frequently P F" using eventually_conj[of P F "\x. \ P x"] by (auto simp add: frequently_def eventually_False) lemma eventually_frequentlyE: assumes "eventually P F" assumes "eventually (\x. \ P x \ Q x) F" "F\bot" shows "frequently Q F" proof - have "eventually Q F" using eventually_conj[OF assms(1,2),simplified] by (auto elim:eventually_mono) then show ?thesis using eventually_frequently[OF \F\bot\] by auto qed lemma eventually_const_iff: "eventually (\x. P) F \ P \ F = bot" by (cases P) (auto simp: eventually_False) lemma eventually_const[simp]: "F \ bot \ eventually (\x. P) F \ P" by (simp add: eventually_const_iff) lemma frequently_const_iff: "frequently (\x. P) F \ P \ F \ bot" by (simp add: frequently_def eventually_const_iff) lemma frequently_const[simp]: "F \ bot \ frequently (\x. P) F \ P" by (simp add: frequently_const_iff) lemma eventually_happens: "eventually P net \ net = bot \ (\x. P x)" by (metis frequentlyE eventually_frequently) lemma eventually_happens': assumes "F \ bot" "eventually P F" shows "\x. P x" using assms eventually_frequently frequentlyE by blast abbreviation (input) trivial_limit :: "'a filter \ bool" where "trivial_limit F \ F = bot" lemma trivial_limit_def: "trivial_limit F \ eventually (\x. False) F" by (rule eventually_False [symmetric]) lemma False_imp_not_eventually: "(\x. \ P x ) \ \ trivial_limit net \ \ eventually (\x. P x) net" by (simp add: eventually_False) lemma eventually_Inf: "eventually P (Inf B) \ (\X\B. finite X \ eventually P (Inf X))" proof - let ?F = "\P. \X\B. finite X \ eventually P (Inf X)" { fix P have "eventually P (Abs_filter ?F) \ ?F P" proof (rule eventually_Abs_filter is_filter.intro)+ show "?F (\x. True)" by (rule exI[of _ "{}"]) (simp add: le_fun_def) next fix P Q assume "?F P" then guess X .. moreover assume "?F Q" then guess Y .. ultimately show "?F (\x. P x \ Q x)" by (intro exI[of _ "X \ Y"]) (auto simp: Inf_union_distrib eventually_inf) next fix P Q assume "?F P" then guess X .. moreover assume "\x. P x \ Q x" ultimately show "?F Q" by (intro exI[of _ X]) (auto elim: eventually_mono) qed } note eventually_F = this have "Inf B = Abs_filter ?F" proof (intro antisym Inf_greatest) show "Inf B \ Abs_filter ?F" by (auto simp: le_filter_def eventually_F dest: Inf_superset_mono) next fix F assume "F \ B" then show "Abs_filter ?F \ F" by (auto simp add: le_filter_def eventually_F intro!: exI[of _ "{F}"]) qed then show ?thesis by (simp add: eventually_F) qed lemma eventually_INF: "eventually P (\b\B. F b) \ (\X\B. finite X \ eventually P (\b\X. F b))" unfolding eventually_Inf [of P "F`B"] by (metis finite_imageI image_mono finite_subset_image) lemma Inf_filter_not_bot: fixes B :: "'a filter set" shows "(\X. X \ B \ finite X \ Inf X \ bot) \ Inf B \ bot" unfolding trivial_limit_def eventually_Inf[of _ B] bot_bool_def [symmetric] bot_fun_def [symmetric] bot_unique by simp lemma INF_filter_not_bot: fixes F :: "'i \ 'a filter" shows "(\X. X \ B \ finite X \ (\b\X. F b) \ bot) \ (\b\B. F b) \ bot" unfolding trivial_limit_def eventually_INF [of _ _ B] bot_bool_def [symmetric] bot_fun_def [symmetric] bot_unique by simp lemma eventually_Inf_base: assumes "B \ {}" and base: "\F G. F \ B \ G \ B \ \x\B. x \ inf F G" shows "eventually P (Inf B) \ (\b\B. eventually P b)" proof (subst eventually_Inf, safe) fix X assume "finite X" "X \ B" then have "\b\B. \x\X. b \ x" proof induct case empty then show ?case using \B \ {}\ by auto next case (insert x X) then obtain b where "b \ B" "\x. x \ X \ b \ x" by auto with \insert x X \ B\ base[of b x] show ?case by (auto intro: order_trans) qed then obtain b where "b \ B" "b \ Inf X" by (auto simp: le_Inf_iff) then show "eventually P (Inf X) \ Bex B (eventually P)" by (intro bexI[of _ b]) (auto simp: le_filter_def) qed (auto intro!: exI[of _ "{x}" for x]) lemma eventually_INF_base: "B \ {} \ (\a b. a \ B \ b \ B \ \x\B. F x \ inf (F a) (F b)) \ eventually P (\b\B. F b) \ (\b\B. eventually P (F b))" by (subst eventually_Inf_base) auto lemma eventually_INF1: "i \ I \ eventually P (F i) \ eventually P (\i\I. F i)" using filter_leD[OF INF_lower] . lemma eventually_INF_finite: assumes "finite A" shows "eventually P (\ x\A. F x) \ (\Q. (\x\A. eventually (Q x) (F x)) \ (\y. (\x\A. Q x y) \ P y))" using assms proof (induction arbitrary: P rule: finite_induct) case (insert a A P) from insert.hyps have [simp]: "x \ a" if "x \ A" for x using that by auto have "eventually P (\ x\insert a A. F x) \ (\Q R S. eventually Q (F a) \ (( (\x\A. eventually (S x) (F x)) \ (\y. (\x\A. S x y) \ R y)) \ (\x. Q x \ R x \ P x)))" unfolding ex_simps by (simp add: eventually_inf insert.IH) also have "\ \ (\Q. (\x\insert a A. eventually (Q x) (F x)) \ (\y. (\x\insert a A. Q x y) \ P y))" proof (safe, goal_cases) case (1 Q R S) thus ?case using 1 by (intro exI[of _ "S(a := Q)"]) auto next case (2 Q) show ?case by (rule exI[of _ "Q a"], rule exI[of _ "\y. \x\A. Q x y"], rule exI[of _ "Q(a := (\_. True))"]) (use 2 in auto) qed finally show ?case . qed auto subsubsection \Map function for filters\ definition filtermap :: "('a \ 'b) \ 'a filter \ 'b filter" where "filtermap f F = Abs_filter (\P. eventually (\x. P (f x)) F)" lemma eventually_filtermap: "eventually P (filtermap f F) = eventually (\x. P (f x)) F" unfolding filtermap_def apply (rule eventually_Abs_filter) apply (rule is_filter.intro) apply (auto elim!: eventually_rev_mp) done lemma filtermap_ident: "filtermap (\x. x) F = F" by (simp add: filter_eq_iff eventually_filtermap) lemma filtermap_filtermap: "filtermap f (filtermap g F) = filtermap (\x. f (g x)) F" by (simp add: filter_eq_iff eventually_filtermap) lemma filtermap_mono: "F \ F' \ filtermap f F \ filtermap f F'" unfolding le_filter_def eventually_filtermap by simp lemma filtermap_bot [simp]: "filtermap f bot = bot" by (simp add: filter_eq_iff eventually_filtermap) lemma filtermap_bot_iff: "filtermap f F = bot \ F = bot" by (simp add: trivial_limit_def eventually_filtermap) lemma filtermap_sup: "filtermap f (sup F1 F2) = sup (filtermap f F1) (filtermap f F2)" by (simp add: filter_eq_iff eventually_filtermap eventually_sup) lemma filtermap_SUP: "filtermap f (\b\B. F b) = (\b\B. filtermap f (F b))" by (simp add: filter_eq_iff eventually_Sup eventually_filtermap) lemma filtermap_inf: "filtermap f (inf F1 F2) \ inf (filtermap f F1) (filtermap f F2)" by (intro inf_greatest filtermap_mono inf_sup_ord) lemma filtermap_INF: "filtermap f (\b\B. F b) \ (\b\B. filtermap f (F b))" by (rule INF_greatest, rule filtermap_mono, erule INF_lower) subsubsection \Contravariant map function for filters\ definition filtercomap :: "('a \ 'b) \ 'b filter \ 'a filter" where "filtercomap f F = Abs_filter (\P. \Q. eventually Q F \ (\x. Q (f x) \ P x))" lemma eventually_filtercomap: "eventually P (filtercomap f F) \ (\Q. eventually Q F \ (\x. Q (f x) \ P x))" unfolding filtercomap_def proof (intro eventually_Abs_filter, unfold_locales, goal_cases) case 1 show ?case by (auto intro!: exI[of _ "\_. True"]) next case (2 P Q) from 2(1) guess P' by (elim exE conjE) note P' = this from 2(2) guess Q' by (elim exE conjE) note Q' = this show ?case by (rule exI[of _ "\x. P' x \ Q' x"]) (insert P' Q', auto intro!: eventually_conj) next case (3 P Q) thus ?case by blast qed lemma filtercomap_ident: "filtercomap (\x. x) F = F" by (auto simp: filter_eq_iff eventually_filtercomap elim!: eventually_mono) lemma filtercomap_filtercomap: "filtercomap f (filtercomap g F) = filtercomap (\x. g (f x)) F" unfolding filter_eq_iff by (auto simp: eventually_filtercomap) lemma filtercomap_mono: "F \ F' \ filtercomap f F \ filtercomap f F'" by (auto simp: eventually_filtercomap le_filter_def) lemma filtercomap_bot [simp]: "filtercomap f bot = bot" by (auto simp: filter_eq_iff eventually_filtercomap) lemma filtercomap_top [simp]: "filtercomap f top = top" by (auto simp: filter_eq_iff eventually_filtercomap) lemma filtercomap_inf: "filtercomap f (inf F1 F2) = inf (filtercomap f F1) (filtercomap f F2)" unfolding filter_eq_iff proof safe fix P assume "eventually P (filtercomap f (F1 \ F2))" then obtain Q R S where *: "eventually Q F1" "eventually R F2" "\x. Q x \ R x \ S x" "\x. S (f x) \ P x" unfolding eventually_filtercomap eventually_inf by blast from * have "eventually (\x. Q (f x)) (filtercomap f F1)" "eventually (\x. R (f x)) (filtercomap f F2)" by (auto simp: eventually_filtercomap) with * show "eventually P (filtercomap f F1 \ filtercomap f F2)" unfolding eventually_inf by blast next fix P assume "eventually P (inf (filtercomap f F1) (filtercomap f F2))" then obtain Q Q' R R' where *: "eventually Q F1" "eventually R F2" "\x. Q (f x) \ Q' x" "\x. R (f x) \ R' x" "\x. Q' x \ R' x \ P x" unfolding eventually_filtercomap eventually_inf by blast from * have "eventually (\x. Q x \ R x) (F1 \ F2)" by (auto simp: eventually_inf) with * show "eventually P (filtercomap f (F1 \ F2))" by (auto simp: eventually_filtercomap) qed lemma filtercomap_sup: "filtercomap f (sup F1 F2) \ sup (filtercomap f F1) (filtercomap f F2)" by (intro sup_least filtercomap_mono inf_sup_ord) lemma filtercomap_INF: "filtercomap f (\b\B. F b) = (\b\B. filtercomap f (F b))" proof - have *: "filtercomap f (\b\B. F b) = (\b\B. filtercomap f (F b))" if "finite B" for B using that by induction (simp_all add: filtercomap_inf) show ?thesis unfolding filter_eq_iff proof fix P have "eventually P (\b\B. filtercomap f (F b)) \ (\X. (X \ B \ finite X) \ eventually P (\b\X. filtercomap f (F b)))" by (subst eventually_INF) blast also have "\ \ (\X. (X \ B \ finite X) \ eventually P (filtercomap f (\b\X. F b)))" by (rule ex_cong) (simp add: *) also have "\ \ eventually P (filtercomap f (\(F ` B)))" unfolding eventually_filtercomap by (subst eventually_INF) blast finally show "eventually P (filtercomap f (\(F ` B))) = eventually P (\b\B. filtercomap f (F b))" .. qed qed lemma filtercomap_SUP: "filtercomap f (\b\B. F b) \ (\b\B. filtercomap f (F b))" by (intro SUP_least filtercomap_mono SUP_upper) lemma filtermap_le_iff_le_filtercomap: "filtermap f F \ G \ F \ filtercomap f G" unfolding le_filter_def eventually_filtermap eventually_filtercomap using eventually_mono by auto lemma filtercomap_neq_bot: assumes "\P. eventually P F \ \x. P (f x)" shows "filtercomap f F \ bot" using assms by (auto simp: trivial_limit_def eventually_filtercomap) lemma filtercomap_neq_bot_surj: assumes "F \ bot" and "surj f" shows "filtercomap f F \ bot" proof (rule filtercomap_neq_bot) fix P assume *: "eventually P F" show "\x. P (f x)" proof (rule ccontr) assume **: "\(\x. P (f x))" from * have "eventually (\_. False) F" proof eventually_elim case (elim x) from \surj f\ obtain y where "x = f y" by auto with elim and ** show False by auto qed with assms show False by (simp add: trivial_limit_def) qed qed lemma eventually_filtercomapI [intro]: assumes "eventually P F" shows "eventually (\x. P (f x)) (filtercomap f F)" using assms by (auto simp: eventually_filtercomap) lemma filtermap_filtercomap: "filtermap f (filtercomap f F) \ F" by (auto simp: le_filter_def eventually_filtermap eventually_filtercomap) lemma filtercomap_filtermap: "filtercomap f (filtermap f F) \ F" unfolding le_filter_def eventually_filtermap eventually_filtercomap by (auto elim!: eventually_mono) subsubsection \Standard filters\ definition principal :: "'a set \ 'a filter" where "principal S = Abs_filter (\P. \x\S. P x)" lemma eventually_principal: "eventually P (principal S) \ (\x\S. P x)" unfolding principal_def by (rule eventually_Abs_filter, rule is_filter.intro) auto lemma eventually_inf_principal: "eventually P (inf F (principal s)) \ eventually (\x. x \ s \ P x) F" unfolding eventually_inf eventually_principal by (auto elim: eventually_mono) lemma principal_UNIV[simp]: "principal UNIV = top" by (auto simp: filter_eq_iff eventually_principal) lemma principal_empty[simp]: "principal {} = bot" by (auto simp: filter_eq_iff eventually_principal) lemma principal_eq_bot_iff: "principal X = bot \ X = {}" by (auto simp add: filter_eq_iff eventually_principal) lemma principal_le_iff[iff]: "principal A \ principal B \ A \ B" by (auto simp: le_filter_def eventually_principal) lemma le_principal: "F \ principal A \ eventually (\x. x \ A) F" unfolding le_filter_def eventually_principal apply safe apply (erule_tac x="\x. x \ A" in allE) apply (auto elim: eventually_mono) done lemma principal_inject[iff]: "principal A = principal B \ A = B" unfolding eq_iff by simp lemma sup_principal[simp]: "sup (principal A) (principal B) = principal (A \ B)" unfolding filter_eq_iff eventually_sup eventually_principal by auto lemma inf_principal[simp]: "inf (principal A) (principal B) = principal (A \ B)" unfolding filter_eq_iff eventually_inf eventually_principal by (auto intro: exI[of _ "\x. x \ A"] exI[of _ "\x. x \ B"]) lemma SUP_principal[simp]: "(\i\I. principal (A i)) = principal (\i\I. A i)" unfolding filter_eq_iff eventually_Sup by (auto simp: eventually_principal) lemma INF_principal_finite: "finite X \ (\x\X. principal (f x)) = principal (\x\X. f x)" by (induct X rule: finite_induct) auto lemma filtermap_principal[simp]: "filtermap f (principal A) = principal (f ` A)" unfolding filter_eq_iff eventually_filtermap eventually_principal by simp lemma filtercomap_principal[simp]: "filtercomap f (principal A) = principal (f -` A)" unfolding filter_eq_iff eventually_filtercomap eventually_principal by fast subsubsection \Order filters\ definition at_top :: "('a::order) filter" where "at_top = (\k. principal {k ..})" lemma at_top_sub: "at_top = (\k\{c::'a::linorder..}. principal {k ..})" by (auto intro!: INF_eq max.cobounded1 max.cobounded2 simp: at_top_def) lemma eventually_at_top_linorder: "eventually P at_top \ (\N::'a::linorder. \n\N. P n)" unfolding at_top_def by (subst eventually_INF_base) (auto simp: eventually_principal intro: max.cobounded1 max.cobounded2) lemma eventually_filtercomap_at_top_linorder: "eventually P (filtercomap f at_top) \ (\N::'a::linorder. \x. f x \ N \ P x)" by (auto simp: eventually_filtercomap eventually_at_top_linorder) lemma eventually_at_top_linorderI: fixes c::"'a::linorder" assumes "\x. c \ x \ P x" shows "eventually P at_top" using assms by (auto simp: eventually_at_top_linorder) lemma eventually_ge_at_top [simp]: "eventually (\x. (c::_::linorder) \ x) at_top" unfolding eventually_at_top_linorder by auto lemma eventually_at_top_dense: "eventually P at_top \ (\N::'a::{no_top, linorder}. \n>N. P n)" proof - have "eventually P (\k. principal {k <..}) \ (\N::'a. \n>N. P n)" by (subst eventually_INF_base) (auto simp: eventually_principal intro: max.cobounded1 max.cobounded2) also have "(\k. principal {k::'a <..}) = at_top" unfolding at_top_def by (intro INF_eq) (auto intro: less_imp_le simp: Ici_subset_Ioi_iff gt_ex) finally show ?thesis . qed lemma eventually_filtercomap_at_top_dense: "eventually P (filtercomap f at_top) \ (\N::'a::{no_top, linorder}. \x. f x > N \ P x)" by (auto simp: eventually_filtercomap eventually_at_top_dense) lemma eventually_at_top_not_equal [simp]: "eventually (\x::'a::{no_top, linorder}. x \ c) at_top" unfolding eventually_at_top_dense by auto lemma eventually_gt_at_top [simp]: "eventually (\x. (c::_::{no_top, linorder}) < x) at_top" unfolding eventually_at_top_dense by auto lemma eventually_all_ge_at_top: assumes "eventually P (at_top :: ('a :: linorder) filter)" shows "eventually (\x. \y\x. P y) at_top" proof - from assms obtain x where "\y. y \ x \ P y" by (auto simp: eventually_at_top_linorder) hence "\z\y. P z" if "y \ x" for y using that by simp thus ?thesis by (auto simp: eventually_at_top_linorder) qed definition at_bot :: "('a::order) filter" where "at_bot = (\k. principal {.. k})" lemma at_bot_sub: "at_bot = (\k\{.. c::'a::linorder}. principal {.. k})" by (auto intro!: INF_eq min.cobounded1 min.cobounded2 simp: at_bot_def) lemma eventually_at_bot_linorder: fixes P :: "'a::linorder \ bool" shows "eventually P at_bot \ (\N. \n\N. P n)" unfolding at_bot_def by (subst eventually_INF_base) (auto simp: eventually_principal intro: min.cobounded1 min.cobounded2) lemma eventually_filtercomap_at_bot_linorder: "eventually P (filtercomap f at_bot) \ (\N::'a::linorder. \x. f x \ N \ P x)" by (auto simp: eventually_filtercomap eventually_at_bot_linorder) lemma eventually_le_at_bot [simp]: "eventually (\x. x \ (c::_::linorder)) at_bot" unfolding eventually_at_bot_linorder by auto lemma eventually_at_bot_dense: "eventually P at_bot \ (\N::'a::{no_bot, linorder}. \nk. principal {..< k}) \ (\N::'a. \nk. principal {..< k::'a}) = at_bot" unfolding at_bot_def by (intro INF_eq) (auto intro: less_imp_le simp: Iic_subset_Iio_iff lt_ex) finally show ?thesis . qed lemma eventually_filtercomap_at_bot_dense: "eventually P (filtercomap f at_bot) \ (\N::'a::{no_bot, linorder}. \x. f x < N \ P x)" by (auto simp: eventually_filtercomap eventually_at_bot_dense) lemma eventually_at_bot_not_equal [simp]: "eventually (\x::'a::{no_bot, linorder}. x \ c) at_bot" unfolding eventually_at_bot_dense by auto lemma eventually_gt_at_bot [simp]: "eventually (\x. x < (c::_::unbounded_dense_linorder)) at_bot" unfolding eventually_at_bot_dense by auto lemma trivial_limit_at_bot_linorder [simp]: "\ trivial_limit (at_bot ::('a::linorder) filter)" unfolding trivial_limit_def by (metis eventually_at_bot_linorder order_refl) lemma trivial_limit_at_top_linorder [simp]: "\ trivial_limit (at_top ::('a::linorder) filter)" unfolding trivial_limit_def by (metis eventually_at_top_linorder order_refl) subsection \Sequentially\ abbreviation sequentially :: "nat filter" where "sequentially \ at_top" lemma eventually_sequentially: "eventually P sequentially \ (\N. \n\N. P n)" by (rule eventually_at_top_linorder) lemma sequentially_bot [simp, intro]: "sequentially \ bot" unfolding filter_eq_iff eventually_sequentially by auto lemmas trivial_limit_sequentially = sequentially_bot lemma eventually_False_sequentially [simp]: "\ eventually (\n. False) sequentially" by (simp add: eventually_False) lemma le_sequentially: "F \ sequentially \ (\N. eventually (\n. N \ n) F)" by (simp add: at_top_def le_INF_iff le_principal) lemma eventually_sequentiallyI [intro?]: assumes "\x. c \ x \ P x" shows "eventually P sequentially" using assms by (auto simp: eventually_sequentially) lemma eventually_sequentially_Suc [simp]: "eventually (\i. P (Suc i)) sequentially \ eventually P sequentially" unfolding eventually_sequentially by (metis Suc_le_D Suc_le_mono le_Suc_eq) lemma eventually_sequentially_seg [simp]: "eventually (\n. P (n + k)) sequentially \ eventually P sequentially" using eventually_sequentially_Suc[of "\n. P (n + k)" for k] by (induction k) auto lemma filtermap_sequentually_ne_bot: "filtermap f sequentially \ bot" by (simp add: filtermap_bot_iff) subsection \Increasing finite subsets\ definition finite_subsets_at_top where "finite_subsets_at_top A = (\ X\{X. finite X \ X \ A}. principal {Y. finite Y \ X \ Y \ Y \ A})" lemma eventually_finite_subsets_at_top: "eventually P (finite_subsets_at_top A) \ (\X. finite X \ X \ A \ (\Y. finite Y \ X \ Y \ Y \ A \ P Y))" unfolding finite_subsets_at_top_def proof (subst eventually_INF_base, goal_cases) show "{X. finite X \ X \ A} \ {}" by auto next case (2 B C) thus ?case by (intro bexI[of _ "B \ C"]) auto qed (simp_all add: eventually_principal) lemma eventually_finite_subsets_at_top_weakI [intro]: assumes "\X. finite X \ X \ A \ P X" shows "eventually P (finite_subsets_at_top A)" proof - have "eventually (\X. finite X \ X \ A) (finite_subsets_at_top A)" by (auto simp: eventually_finite_subsets_at_top) thus ?thesis by eventually_elim (use assms in auto) qed lemma finite_subsets_at_top_neq_bot [simp]: "finite_subsets_at_top A \ bot" proof - have "\eventually (\x. False) (finite_subsets_at_top A)" by (auto simp: eventually_finite_subsets_at_top) thus ?thesis by auto qed lemma filtermap_image_finite_subsets_at_top: assumes "inj_on f A" shows "filtermap ((`) f) (finite_subsets_at_top A) = finite_subsets_at_top (f ` A)" unfolding filter_eq_iff eventually_filtermap proof (safe, goal_cases) case (1 P) then obtain X where X: "finite X" "X \ A" "\Y. finite Y \ X \ Y \ Y \ A \ P (f ` Y)" unfolding eventually_finite_subsets_at_top by force show ?case unfolding eventually_finite_subsets_at_top eventually_filtermap proof (rule exI[of _ "f ` X"], intro conjI allI impI, goal_cases) case (3 Y) with assms and X(1,2) have "P (f ` (f -` Y \ A))" using X(1,2) by (intro X(3) finite_vimage_IntI) auto also have "f ` (f -` Y \ A) = Y" using assms 3 by blast finally show ?case . qed (insert assms X(1,2), auto intro!: finite_vimage_IntI) next case (2 P) then obtain X where X: "finite X" "X \ f ` A" "\Y. finite Y \ X \ Y \ Y \ f ` A \ P Y" unfolding eventually_finite_subsets_at_top by force show ?case unfolding eventually_finite_subsets_at_top eventually_filtermap proof (rule exI[of _ "f -` X \ A"], intro conjI allI impI, goal_cases) case (3 Y) with X(1,2) and assms show ?case by (intro X(3)) force+ qed (insert assms X(1), auto intro!: finite_vimage_IntI) qed lemma eventually_finite_subsets_at_top_finite: assumes "finite A" shows "eventually P (finite_subsets_at_top A) \ P A" unfolding eventually_finite_subsets_at_top using assms by force lemma finite_subsets_at_top_finite: "finite A \ finite_subsets_at_top A = principal {A}" by (auto simp: filter_eq_iff eventually_finite_subsets_at_top_finite eventually_principal) subsection \The cofinite filter\ definition "cofinite = Abs_filter (\P. finite {x. \ P x})" abbreviation Inf_many :: "('a \ bool) \ bool" (binder "\\<^sub>\" 10) where "Inf_many P \ frequently P cofinite" abbreviation Alm_all :: "('a \ bool) \ bool" (binder "\\<^sub>\" 10) where "Alm_all P \ eventually P cofinite" notation (ASCII) Inf_many (binder "INFM " 10) and Alm_all (binder "MOST " 10) lemma eventually_cofinite: "eventually P cofinite \ finite {x. \ P x}" unfolding cofinite_def proof (rule eventually_Abs_filter, rule is_filter.intro) fix P Q :: "'a \ bool" assume "finite {x. \ P x}" "finite {x. \ Q x}" from finite_UnI[OF this] show "finite {x. \ (P x \ Q x)}" by (rule rev_finite_subset) auto next fix P Q :: "'a \ bool" assume P: "finite {x. \ P x}" and *: "\x. P x \ Q x" from * show "finite {x. \ Q x}" by (intro finite_subset[OF _ P]) auto qed simp lemma frequently_cofinite: "frequently P cofinite \ \ finite {x. P x}" by (simp add: frequently_def eventually_cofinite) lemma cofinite_bot[simp]: "cofinite = (bot::'a filter) \ finite (UNIV :: 'a set)" unfolding trivial_limit_def eventually_cofinite by simp lemma cofinite_eq_sequentially: "cofinite = sequentially" unfolding filter_eq_iff eventually_sequentially eventually_cofinite proof safe fix P :: "nat \ bool" assume [simp]: "finite {x. \ P x}" show "\N. \n\N. P n" proof cases assume "{x. \ P x} \ {}" then show ?thesis by (intro exI[of _ "Suc (Max {x. \ P x})"]) (auto simp: Suc_le_eq) qed auto next fix P :: "nat \ bool" and N :: nat assume "\n\N. P n" then have "{x. \ P x} \ {..< N}" by (auto simp: not_le) then show "finite {x. \ P x}" by (blast intro: finite_subset) qed subsubsection \Product of filters\ definition prod_filter :: "'a filter \ 'b filter \ ('a \ 'b) filter" (infixr "\\<^sub>F" 80) where "prod_filter F G = (\(P, Q)\{(P, Q). eventually P F \ eventually Q G}. principal {(x, y). P x \ Q y})" lemma eventually_prod_filter: "eventually P (F \\<^sub>F G) \ (\Pf Pg. eventually Pf F \ eventually Pg G \ (\x y. Pf x \ Pg y \ P (x, y)))" unfolding prod_filter_def proof (subst eventually_INF_base, goal_cases) case 2 moreover have "eventually Pf F \ eventually Qf F \ eventually Pg G \ eventually Qg G \ \P Q. eventually P F \ eventually Q G \ Collect P \ Collect Q \ Collect Pf \ Collect Pg \ Collect Qf \ Collect Qg" for Pf Pg Qf Qg by (intro conjI exI[of _ "inf Pf Qf"] exI[of _ "inf Pg Qg"]) (auto simp: inf_fun_def eventually_conj) ultimately show ?case by auto qed (auto simp: eventually_principal intro: eventually_True) lemma eventually_prod1: assumes "B \ bot" shows "(\\<^sub>F (x, y) in A \\<^sub>F B. P x) \ (\\<^sub>F x in A. P x)" unfolding eventually_prod_filter proof safe fix R Q assume *: "\\<^sub>F x in A. R x" "\\<^sub>F x in B. Q x" "\x y. R x \ Q y \ P x" with \B \ bot\ obtain y where "Q y" by (auto dest: eventually_happens) with * show "eventually P A" by (force elim: eventually_mono) next assume "eventually P A" then show "\Pf Pg. eventually Pf A \ eventually Pg B \ (\x y. Pf x \ Pg y \ P x)" by (intro exI[of _ P] exI[of _ "\x. True"]) auto qed lemma eventually_prod2: assumes "A \ bot" shows "(\\<^sub>F (x, y) in A \\<^sub>F B. P y) \ (\\<^sub>F y in B. P y)" unfolding eventually_prod_filter proof safe fix R Q assume *: "\\<^sub>F x in A. R x" "\\<^sub>F x in B. Q x" "\x y. R x \ Q y \ P y" with \A \ bot\ obtain x where "R x" by (auto dest: eventually_happens) with * show "eventually P B" by (force elim: eventually_mono) next assume "eventually P B" then show "\Pf Pg. eventually Pf A \ eventually Pg B \ (\x y. Pf x \ Pg y \ P y)" by (intro exI[of _ P] exI[of _ "\x. True"]) auto qed lemma INF_filter_bot_base: fixes F :: "'a \ 'b filter" assumes *: "\i j. i \ I \ j \ I \ \k\I. F k \ F i \ F j" shows "(\i\I. F i) = bot \ (\i\I. F i = bot)" proof (cases "\i\I. F i = bot") case True then have "(\i\I. F i) \ bot" by (auto intro: INF_lower2) with True show ?thesis by (auto simp: bot_unique) next case False moreover have "(\i\I. F i) \ bot" proof (cases "I = {}") case True then show ?thesis by (auto simp add: filter_eq_iff) next case False': False show ?thesis proof (rule INF_filter_not_bot) fix J assume "finite J" "J \ I" then have "\k\I. F k \ (\i\J. F i)" proof (induct J) case empty then show ?case using \I \ {}\ by auto next case (insert i J) then obtain k where "k \ I" "F k \ (\i\J. F i)" by auto with insert *[of i k] show ?case by auto qed with False show "(\i\J. F i) \ \" by (auto simp: bot_unique) qed qed ultimately show ?thesis by auto qed lemma Collect_empty_eq_bot: "Collect P = {} \ P = \" by auto lemma prod_filter_eq_bot: "A \\<^sub>F B = bot \ A = bot \ B = bot" unfolding trivial_limit_def proof assume "\\<^sub>F x in A \\<^sub>F B. False" then obtain Pf Pg where Pf: "eventually (\x. Pf x) A" and Pg: "eventually (\y. Pg y) B" and *: "\x y. Pf x \ Pg y \ False" unfolding eventually_prod_filter by fast from * have "(\x. \ Pf x) \ (\y. \ Pg y)" by fast with Pf Pg show "(\\<^sub>F x in A. False) \ (\\<^sub>F x in B. False)" by auto next assume "(\\<^sub>F x in A. False) \ (\\<^sub>F x in B. False)" then show "\\<^sub>F x in A \\<^sub>F B. False" unfolding eventually_prod_filter by (force intro: eventually_True) qed lemma prod_filter_mono: "F \ F' \ G \ G' \ F \\<^sub>F G \ F' \\<^sub>F G'" by (auto simp: le_filter_def eventually_prod_filter) lemma prod_filter_mono_iff: assumes nAB: "A \ bot" "B \ bot" shows "A \\<^sub>F B \ C \\<^sub>F D \ A \ C \ B \ D" proof safe assume *: "A \\<^sub>F B \ C \\<^sub>F D" with assms have "A \\<^sub>F B \ bot" by (auto simp: bot_unique prod_filter_eq_bot) with * have "C \\<^sub>F D \ bot" by (auto simp: bot_unique) then have nCD: "C \ bot" "D \ bot" by (auto simp: prod_filter_eq_bot) show "A \ C" proof (rule filter_leI) fix P assume "eventually P C" with *[THEN filter_leD, of "\(x, y). P x"] show "eventually P A" using nAB nCD by (simp add: eventually_prod1 eventually_prod2) qed show "B \ D" proof (rule filter_leI) fix P assume "eventually P D" with *[THEN filter_leD, of "\(x, y). P y"] show "eventually P B" using nAB nCD by (simp add: eventually_prod1 eventually_prod2) qed qed (intro prod_filter_mono) lemma eventually_prod_same: "eventually P (F \\<^sub>F F) \ (\Q. eventually Q F \ (\x y. Q x \ Q y \ P (x, y)))" unfolding eventually_prod_filter apply safe apply (rule_tac x="inf Pf Pg" in exI) apply (auto simp: inf_fun_def intro!: eventually_conj) done lemma eventually_prod_sequentially: "eventually P (sequentially \\<^sub>F sequentially) \ (\N. \m \ N. \n \ N. P (n, m))" unfolding eventually_prod_same eventually_sequentially by auto lemma principal_prod_principal: "principal A \\<^sub>F principal B = principal (A \ B)" unfolding filter_eq_iff eventually_prod_filter eventually_principal by (fast intro: exI[of _ "\x. x \ A"] exI[of _ "\x. x \ B"]) lemma le_prod_filterI: "filtermap fst F \ A \ filtermap snd F \ B \ F \ A \\<^sub>F B" unfolding le_filter_def eventually_filtermap eventually_prod_filter by (force elim: eventually_elim2) lemma filtermap_fst_prod_filter: "filtermap fst (A \\<^sub>F B) \ A" unfolding le_filter_def eventually_filtermap eventually_prod_filter by (force intro: eventually_True) lemma filtermap_snd_prod_filter: "filtermap snd (A \\<^sub>F B) \ B" unfolding le_filter_def eventually_filtermap eventually_prod_filter by (force intro: eventually_True) lemma prod_filter_INF: assumes "I \ {}" and "J \ {}" shows "(\i\I. A i) \\<^sub>F (\j\J. B j) = (\i\I. \j\J. A i \\<^sub>F B j)" proof (rule antisym) from \I \ {}\ obtain i where "i \ I" by auto from \J \ {}\ obtain j where "j \ J" by auto show "(\i\I. \j\J. A i \\<^sub>F B j) \ (\i\I. A i) \\<^sub>F (\j\J. B j)" by (fast intro: le_prod_filterI INF_greatest INF_lower2 order_trans[OF filtermap_INF] \i \ I\ \j \ J\ filtermap_fst_prod_filter filtermap_snd_prod_filter) show "(\i\I. A i) \\<^sub>F (\j\J. B j) \ (\i\I. \j\J. A i \\<^sub>F B j)" by (intro INF_greatest prod_filter_mono INF_lower) qed lemma filtermap_Pair: "filtermap (\x. (f x, g x)) F \ filtermap f F \\<^sub>F filtermap g F" by (rule le_prod_filterI, simp_all add: filtermap_filtermap) lemma eventually_prodI: "eventually P F \ eventually Q G \ eventually (\x. P (fst x) \ Q (snd x)) (F \\<^sub>F G)" unfolding eventually_prod_filter by auto lemma prod_filter_INF1: "I \ {} \ (\i\I. A i) \\<^sub>F B = (\i\I. A i \\<^sub>F B)" using prod_filter_INF[of I "{B}" A "\x. x"] by simp lemma prod_filter_INF2: "J \ {} \ A \\<^sub>F (\i\J. B i) = (\i\J. A \\<^sub>F B i)" using prod_filter_INF[of "{A}" J "\x. x" B] by simp lemma prod_filtermap1: "prod_filter (filtermap f F) G = filtermap (apfst f) (prod_filter F G)" apply(clarsimp simp add: filter_eq_iff eventually_filtermap eventually_prod_filter; safe) subgoal by auto subgoal for P Q R by(rule exI[where x="\y. \x. y = f x \ Q x"])(auto intro: eventually_mono) done lemma prod_filtermap2: "prod_filter F (filtermap g G) = filtermap (apsnd g) (prod_filter F G)" apply(clarsimp simp add: filter_eq_iff eventually_filtermap eventually_prod_filter; safe) subgoal by auto subgoal for P Q R by(auto intro: exI[where x="\y. \x. y = g x \ R x"] eventually_mono) done lemma prod_filter_assoc: "prod_filter (prod_filter F G) H = filtermap (\(x, y, z). ((x, y), z)) (prod_filter F (prod_filter G H))" apply(clarsimp simp add: filter_eq_iff eventually_filtermap eventually_prod_filter; safe) subgoal for P Q R S T by(auto 4 4 intro: exI[where x="\(a, b). T a \ S b"]) subgoal for P Q R S T by(auto 4 3 intro: exI[where x="\(a, b). Q a \ S b"]) done lemma prod_filter_principal_singleton: "prod_filter (principal {x}) F = filtermap (Pair x) F" by(fastforce simp add: filter_eq_iff eventually_prod_filter eventually_principal eventually_filtermap elim: eventually_mono intro: exI[where x="\a. a = x"]) lemma prod_filter_principal_singleton2: "prod_filter F (principal {x}) = filtermap (\a. (a, x)) F" by(fastforce simp add: filter_eq_iff eventually_prod_filter eventually_principal eventually_filtermap elim: eventually_mono intro: exI[where x="\a. a = x"]) lemma prod_filter_commute: "prod_filter F G = filtermap prod.swap (prod_filter G F)" by(auto simp add: filter_eq_iff eventually_prod_filter eventually_filtermap) subsection \Limits\ definition filterlim :: "('a \ 'b) \ 'b filter \ 'a filter \ bool" where "filterlim f F2 F1 \ filtermap f F1 \ F2" syntax "_LIM" :: "pttrns \ 'a \ 'b \ 'a \ bool" ("(3LIM (_)/ (_)./ (_) :> (_))" [1000, 10, 0, 10] 10) translations "LIM x F1. f :> F2" == "CONST filterlim (\x. f) F2 F1" lemma filterlim_top [simp]: "filterlim f top F" by (simp add: filterlim_def) lemma filterlim_iff: "(LIM x F1. f x :> F2) \ (\P. eventually P F2 \ eventually (\x. P (f x)) F1)" unfolding filterlim_def le_filter_def eventually_filtermap .. lemma filterlim_compose: "filterlim g F3 F2 \ filterlim f F2 F1 \ filterlim (\x. g (f x)) F3 F1" unfolding filterlim_def filtermap_filtermap[symmetric] by (metis filtermap_mono order_trans) lemma filterlim_mono: "filterlim f F2 F1 \ F2 \ F2' \ F1' \ F1 \ filterlim f F2' F1'" unfolding filterlim_def by (metis filtermap_mono order_trans) lemma filterlim_ident: "LIM x F. x :> F" by (simp add: filterlim_def filtermap_ident) lemma filterlim_cong: "F1 = F1' \ F2 = F2' \ eventually (\x. f x = g x) F2 \ filterlim f F1 F2 = filterlim g F1' F2'" by (auto simp: filterlim_def le_filter_def eventually_filtermap elim: eventually_elim2) lemma filterlim_mono_eventually: assumes "filterlim f F G" and ord: "F \ F'" "G' \ G" assumes eq: "eventually (\x. f x = f' x) G'" shows "filterlim f' F' G'" apply (rule filterlim_cong[OF refl refl eq, THEN iffD1]) apply (rule filterlim_mono[OF _ ord]) apply fact done lemma filtermap_mono_strong: "inj f \ filtermap f F \ filtermap f G \ F \ G" apply (safe intro!: filtermap_mono) apply (auto simp: le_filter_def eventually_filtermap) apply (erule_tac x="\x. P (inv f x)" in allE) apply auto done lemma eventually_compose_filterlim: assumes "eventually P F" "filterlim f F G" shows "eventually (\x. P (f x)) G" using assms by (simp add: filterlim_iff) lemma filtermap_eq_strong: "inj f \ filtermap f F = filtermap f G \ F = G" by (simp add: filtermap_mono_strong eq_iff) lemma filtermap_fun_inverse: assumes g: "filterlim g F G" assumes f: "filterlim f G F" assumes ev: "eventually (\x. f (g x) = x) G" shows "filtermap f F = G" proof (rule antisym) show "filtermap f F \ G" using f unfolding filterlim_def . have "G = filtermap f (filtermap g G)" using ev by (auto elim: eventually_elim2 simp: filter_eq_iff eventually_filtermap) also have "\ \ filtermap f F" using g by (intro filtermap_mono) (simp add: filterlim_def) finally show "G \ filtermap f F" . qed lemma filterlim_principal: "(LIM x F. f x :> principal S) \ (eventually (\x. f x \ S) F)" unfolding filterlim_def eventually_filtermap le_principal .. lemma filterlim_filtercomap [intro]: "filterlim f F (filtercomap f F)" unfolding filterlim_def by (rule filtermap_filtercomap) lemma filterlim_inf: "(LIM x F1. f x :> inf F2 F3) \ ((LIM x F1. f x :> F2) \ (LIM x F1. f x :> F3))" unfolding filterlim_def by simp lemma filterlim_INF: "(LIM x F. f x :> (\b\B. G b)) \ (\b\B. LIM x F. f x :> G b)" unfolding filterlim_def le_INF_iff .. lemma filterlim_INF_INF: "(\m. m \ J \ \i\I. filtermap f (F i) \ G m) \ LIM x (\i\I. F i). f x :> (\j\J. G j)" unfolding filterlim_def by (rule order_trans[OF filtermap_INF INF_mono]) lemma filterlim_INF': "x \ A \ filterlim f F (G x) \ filterlim f F (\ x\A. G x)" unfolding filterlim_def by (rule order.trans[OF filtermap_mono[OF INF_lower]]) lemma filterlim_filtercomap_iff: "filterlim f (filtercomap g G) F \ filterlim (g \ f) G F" by (simp add: filterlim_def filtermap_le_iff_le_filtercomap filtercomap_filtercomap o_def) lemma filterlim_iff_le_filtercomap: "filterlim f F G \ G \ filtercomap f F" by (simp add: filterlim_def filtermap_le_iff_le_filtercomap) lemma filterlim_base: "(\m x. m \ J \ i m \ I) \ (\m x. m \ J \ x \ F (i m) \ f x \ G m) \ LIM x (\i\I. principal (F i)). f x :> (\j\J. principal (G j))" by (force intro!: filterlim_INF_INF simp: image_subset_iff) lemma filterlim_base_iff: assumes "I \ {}" and chain: "\i j. i \ I \ j \ I \ F i \ F j \ F j \ F i" shows "(LIM x (\i\I. principal (F i)). f x :> \j\J. principal (G j)) \ (\j\J. \i\I. \x\F i. f x \ G j)" unfolding filterlim_INF filterlim_principal proof (subst eventually_INF_base) fix i j assume "i \ I" "j \ I" with chain[OF this] show "\x\I. principal (F x) \ inf (principal (F i)) (principal (F j))" by auto qed (auto simp: eventually_principal \I \ {}\) lemma filterlim_filtermap: "filterlim f F1 (filtermap g F2) = filterlim (\x. f (g x)) F1 F2" unfolding filterlim_def filtermap_filtermap .. lemma filterlim_sup: "filterlim f F F1 \ filterlim f F F2 \ filterlim f F (sup F1 F2)" unfolding filterlim_def filtermap_sup by auto lemma filterlim_sequentially_Suc: "(LIM x sequentially. f (Suc x) :> F) \ (LIM x sequentially. f x :> F)" unfolding filterlim_iff by (subst eventually_sequentially_Suc) simp lemma filterlim_Suc: "filterlim Suc sequentially sequentially" by (simp add: filterlim_iff eventually_sequentially) lemma filterlim_If: "LIM x inf F (principal {x. P x}). f x :> G \ LIM x inf F (principal {x. \ P x}). g x :> G \ LIM x F. if P x then f x else g x :> G" unfolding filterlim_iff eventually_inf_principal by (auto simp: eventually_conj_iff) lemma filterlim_Pair: "LIM x F. f x :> G \ LIM x F. g x :> H \ LIM x F. (f x, g x) :> G \\<^sub>F H" unfolding filterlim_def by (rule order_trans[OF filtermap_Pair prod_filter_mono]) subsection \Limits to \<^const>\at_top\ and \<^const>\at_bot\\ lemma filterlim_at_top: fixes f :: "'a \ ('b::linorder)" shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z \ f x) F)" by (auto simp: filterlim_iff eventually_at_top_linorder elim!: eventually_mono) lemma filterlim_at_top_mono: "LIM x F. f x :> at_top \ eventually (\x. f x \ (g x::'a::linorder)) F \ LIM x F. g x :> at_top" by (auto simp: filterlim_at_top elim: eventually_elim2 intro: order_trans) lemma filterlim_at_top_dense: fixes f :: "'a \ ('b::unbounded_dense_linorder)" shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z < f x) F)" by (metis eventually_mono[of _ F] eventually_gt_at_top order_less_imp_le filterlim_at_top[of f F] filterlim_iff[of f at_top F]) lemma filterlim_at_top_ge: fixes f :: "'a \ ('b::linorder)" and c :: "'b" shows "(LIM x F. f x :> at_top) \ (\Z\c. eventually (\x. Z \ f x) F)" unfolding at_top_sub[of c] filterlim_INF by (auto simp add: filterlim_principal) lemma filterlim_at_top_at_top: fixes f :: "'a::linorder \ 'b::linorder" assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" assumes Q: "eventually Q at_top" assumes P: "eventually P at_top" shows "filterlim f at_top at_top" proof - from P obtain x where x: "\y. x \ y \ P y" unfolding eventually_at_top_linorder by auto show ?thesis proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) fix z assume "x \ z" with x have "P z" by auto have "eventually (\x. g z \ x) at_top" by (rule eventually_ge_at_top) with Q show "eventually (\x. z \ f x) at_top" by eventually_elim (metis mono bij \P z\) qed qed lemma filterlim_at_top_gt: fixes f :: "'a \ ('b::unbounded_dense_linorder)" and c :: "'b" shows "(LIM x F. f x :> at_top) \ (\Z>c. eventually (\x. Z \ f x) F)" by (metis filterlim_at_top order_less_le_trans gt_ex filterlim_at_top_ge) lemma filterlim_at_bot: fixes f :: "'a \ ('b::linorder)" shows "(LIM x F. f x :> at_bot) \ (\Z. eventually (\x. f x \ Z) F)" by (auto simp: filterlim_iff eventually_at_bot_linorder elim!: eventually_mono) lemma filterlim_at_bot_dense: fixes f :: "'a \ ('b::{dense_linorder, no_bot})" shows "(LIM x F. f x :> at_bot) \ (\Z. eventually (\x. f x < Z) F)" proof (auto simp add: filterlim_at_bot[of f F]) fix Z :: 'b from lt_ex [of Z] obtain Z' where 1: "Z' < Z" .. assume "\Z. eventually (\x. f x \ Z) F" hence "eventually (\x. f x \ Z') F" by auto thus "eventually (\x. f x < Z) F" apply (rule eventually_mono) using 1 by auto next fix Z :: 'b show "\Z. eventually (\x. f x < Z) F \ eventually (\x. f x \ Z) F" by (drule spec [of _ Z], erule eventually_mono, auto simp add: less_imp_le) qed lemma filterlim_at_bot_le: fixes f :: "'a \ ('b::linorder)" and c :: "'b" shows "(LIM x F. f x :> at_bot) \ (\Z\c. eventually (\x. Z \ f x) F)" unfolding filterlim_at_bot proof safe fix Z assume *: "\Z\c. eventually (\x. Z \ f x) F" with *[THEN spec, of "min Z c"] show "eventually (\x. Z \ f x) F" by (auto elim!: eventually_mono) qed simp lemma filterlim_at_bot_lt: fixes f :: "'a \ ('b::unbounded_dense_linorder)" and c :: "'b" shows "(LIM x F. f x :> at_bot) \ (\Zx. Z \ f x) F)" by (metis filterlim_at_bot filterlim_at_bot_le lt_ex order_le_less_trans) lemma filterlim_finite_subsets_at_top: "filterlim f (finite_subsets_at_top A) F \ (\X. finite X \ X \ A \ eventually (\y. finite (f y) \ X \ f y \ f y \ A) F)" (is "?lhs = ?rhs") proof assume ?lhs thus ?rhs proof (safe, goal_cases) case (1 X) hence *: "(\\<^sub>F x in F. P (f x))" if "eventually P (finite_subsets_at_top A)" for P using that by (auto simp: filterlim_def le_filter_def eventually_filtermap) have "\\<^sub>F Y in finite_subsets_at_top A. finite Y \ X \ Y \ Y \ A" using 1 unfolding eventually_finite_subsets_at_top by force thus ?case by (intro *) auto qed next assume rhs: ?rhs show ?lhs unfolding filterlim_def le_filter_def eventually_finite_subsets_at_top proof (safe, goal_cases) case (1 P X) with rhs have "\\<^sub>F y in F. finite (f y) \ X \ f y \ f y \ A" by auto thus "eventually P (filtermap f F)" unfolding eventually_filtermap by eventually_elim (insert 1, auto) qed qed lemma filterlim_atMost_at_top: "filterlim (\n. {..n}) (finite_subsets_at_top (UNIV :: nat set)) at_top" unfolding filterlim_finite_subsets_at_top proof (safe, goal_cases) case (1 X) then obtain n where n: "X \ {..n}" by (auto simp: finite_nat_set_iff_bounded_le) show ?case using eventually_ge_at_top[of n] by eventually_elim (insert n, auto) qed lemma filterlim_lessThan_at_top: "filterlim (\n. {.. {..Setup \<^typ>\'a filter\ for lifting and transfer\ lemma filtermap_id [simp, id_simps]: "filtermap id = id" by(simp add: fun_eq_iff id_def filtermap_ident) lemma filtermap_id' [simp]: "filtermap (\x. x) = (\F. F)" using filtermap_id unfolding id_def . context includes lifting_syntax begin definition map_filter_on :: "'a set \ ('a \ 'b) \ 'a filter \ 'b filter" where "map_filter_on X f F = Abs_filter (\P. eventually (\x. P (f x) \ x \ X) F)" lemma is_filter_map_filter_on: "is_filter (\P. \\<^sub>F x in F. P (f x) \ x \ X) \ eventually (\x. x \ X) F" proof(rule iffI; unfold_locales) show "\\<^sub>F x in F. True \ x \ X" if "eventually (\x. x \ X) F" using that by simp show "\\<^sub>F x in F. (P (f x) \ Q (f x)) \ x \ X" if "\\<^sub>F x in F. P (f x) \ x \ X" "\\<^sub>F x in F. Q (f x) \ x \ X" for P Q using eventually_conj[OF that] by(auto simp add: conj_ac cong: conj_cong) show "\\<^sub>F x in F. Q (f x) \ x \ X" if "\x. P x \ Q x" "\\<^sub>F x in F. P (f x) \ x \ X" for P Q using that(2) by(rule eventually_mono)(use that(1) in auto) show "eventually (\x. x \ X) F" if "is_filter (\P. \\<^sub>F x in F. P (f x) \ x \ X)" using is_filter.True[OF that] by simp qed lemma eventually_map_filter_on: "eventually P (map_filter_on X f F) = (\\<^sub>F x in F. P (f x) \ x \ X)" if "eventually (\x. x \ X) F" by(simp add: is_filter_map_filter_on map_filter_on_def eventually_Abs_filter that) lemma map_filter_on_UNIV: "map_filter_on UNIV = filtermap" by(simp add: map_filter_on_def filtermap_def fun_eq_iff) lemma map_filter_on_comp: "map_filter_on X f (map_filter_on Y g F) = map_filter_on Y (f \ g) F" if "g ` Y \ X" and "eventually (\x. x \ Y) F" unfolding map_filter_on_def using that(1) by(auto simp add: eventually_Abs_filter that(2) is_filter_map_filter_on intro!: arg_cong[where f=Abs_filter] arg_cong2[where f=eventually]) inductive rel_filter :: "('a \ 'b \ bool) \ 'a filter \ 'b filter \ bool" for R F G where "rel_filter R F G" if "eventually (case_prod R) Z" "map_filter_on {(x, y). R x y} fst Z = F" "map_filter_on {(x, y). R x y} snd Z = G" lemma rel_filter_eq [relator_eq]: "rel_filter (=) = (=)" proof(intro ext iffI)+ show "F = G" if "rel_filter (=) F G" for F G using that by cases(clarsimp simp add: filter_eq_iff eventually_map_filter_on split_def cong: rev_conj_cong) show "rel_filter (=) F G" if "F = G" for F G unfolding \F = G\ proof let ?Z = "map_filter_on UNIV (\x. (x, x)) G" have [simp]: "range (\x. (x, x)) \ {(x, y). x = y}" by auto show "map_filter_on {(x, y). x = y} fst ?Z = G" and "map_filter_on {(x, y). x = y} snd ?Z = G" by(simp_all add: map_filter_on_comp)(simp_all add: map_filter_on_UNIV o_def) show "\\<^sub>F (x, y) in ?Z. x = y" by(simp add: eventually_map_filter_on) qed qed lemma rel_filter_mono [relator_mono]: "rel_filter A \ rel_filter B" if le: "A \ B" proof(clarify elim!: rel_filter.cases) show "rel_filter B (map_filter_on {(x, y). A x y} fst Z) (map_filter_on {(x, y). A x y} snd Z)" (is "rel_filter _ ?X ?Y") if "\\<^sub>F (x, y) in Z. A x y" for Z proof let ?Z = "map_filter_on {(x, y). A x y} id Z" show "\\<^sub>F (x, y) in ?Z. B x y" using le that by(simp add: eventually_map_filter_on le_fun_def split_def conj_commute cong: conj_cong) have [simp]: "{(x, y). A x y} \ {(x, y). B x y}" using le by auto show "map_filter_on {(x, y). B x y} fst ?Z = ?X" "map_filter_on {(x, y). B x y} snd ?Z = ?Y" using le that by(simp_all add: le_fun_def map_filter_on_comp) qed qed lemma rel_filter_conversep: "rel_filter A\\ = (rel_filter A)\\" proof(safe intro!: ext elim!: rel_filter.cases) show *: "rel_filter A (map_filter_on {(x, y). A\\ x y} snd Z) (map_filter_on {(x, y). A\\ x y} fst Z)" (is "rel_filter _ ?X ?Y") if "\\<^sub>F (x, y) in Z. A\\ x y" for A Z proof let ?Z = "map_filter_on {(x, y). A y x} prod.swap Z" show "\\<^sub>F (x, y) in ?Z. A x y" using that by(simp add: eventually_map_filter_on) have [simp]: "prod.swap ` {(x, y). A y x} \ {(x, y). A x y}" by auto show "map_filter_on {(x, y). A x y} fst ?Z = ?X" "map_filter_on {(x, y). A x y} snd ?Z = ?Y" using that by(simp_all add: map_filter_on_comp o_def) qed show "rel_filter A\\ (map_filter_on {(x, y). A x y} snd Z) (map_filter_on {(x, y). A x y} fst Z)" if "\\<^sub>F (x, y) in Z. A x y" for Z using *[of "A\\" Z] that by simp qed lemma rel_filter_distr [relator_distr]: "rel_filter A OO rel_filter B = rel_filter (A OO B)" proof(safe intro!: ext elim!: rel_filter.cases) let ?AB = "{(x, y). (A OO B) x y}" show "(rel_filter A OO rel_filter B) (map_filter_on {(x, y). (A OO B) x y} fst Z) (map_filter_on {(x, y). (A OO B) x y} snd Z)" (is "(_ OO _) ?F ?H") if "\\<^sub>F (x, y) in Z. (A OO B) x y" for Z proof let ?G = "map_filter_on ?AB (\(x, y). SOME z. A x z \ B z y) Z" show "rel_filter A ?F ?G" proof let ?Z = "map_filter_on ?AB (\(x, y). (x, SOME z. A x z \ B z y)) Z" show "\\<^sub>F (x, y) in ?Z. A x y" using that by(auto simp add: eventually_map_filter_on split_def elim!: eventually_mono intro: someI2) have [simp]: "(\p. (fst p, SOME z. A (fst p) z \ B z (snd p))) ` {p. (A OO B) (fst p) (snd p)} \ {p. A (fst p) (snd p)}" by(auto intro: someI2) show "map_filter_on {(x, y). A x y} fst ?Z = ?F" "map_filter_on {(x, y). A x y} snd ?Z = ?G" using that by(simp_all add: map_filter_on_comp split_def o_def) qed show "rel_filter B ?G ?H" proof let ?Z = "map_filter_on ?AB (\(x, y). (SOME z. A x z \ B z y, y)) Z" show "\\<^sub>F (x, y) in ?Z. B x y" using that by(auto simp add: eventually_map_filter_on split_def elim!: eventually_mono intro: someI2) have [simp]: "(\p. (SOME z. A (fst p) z \ B z (snd p), snd p)) ` {p. (A OO B) (fst p) (snd p)} \ {p. B (fst p) (snd p)}" by(auto intro: someI2) show "map_filter_on {(x, y). B x y} fst ?Z = ?G" "map_filter_on {(x, y). B x y} snd ?Z = ?H" using that by(simp_all add: map_filter_on_comp split_def o_def) qed qed fix F G assume F: "\\<^sub>F (x, y) in F. A x y" and G: "\\<^sub>F (x, y) in G. B x y" and eq: "map_filter_on {(x, y). B x y} fst G = map_filter_on {(x, y). A x y} snd F" (is "?Y2 = ?Y1") let ?X = "map_filter_on {(x, y). A x y} fst F" and ?Z = "(map_filter_on {(x, y). B x y} snd G)" have step: "\P'\P. \Q' \ Q. eventually P' F \ eventually Q' G \ {y. \x. P' (x, y)} = {y. \z. Q' (y, z)}" if P: "eventually P F" and Q: "eventually Q G" for P Q proof - let ?P = "\(x, y). P (x, y) \ A x y" and ?Q = "\(y, z). Q (y, z) \ B y z" define P' where "P' \ \(x, y). ?P (x, y) \ (\z. ?Q (y, z))" define Q' where "Q' \ \(y, z). ?Q (y, z) \ (\x. ?P (x, y))" have "P' \ P" "Q' \ Q" "{y. \x. P' (x, y)} = {y. \z. Q' (y, z)}" by(auto simp add: P'_def Q'_def) moreover from P Q F G have P': "eventually ?P F" and Q': "eventually ?Q G" by(simp_all add: eventually_conj_iff split_def) from P' F have "\\<^sub>F y in ?Y1. \x. P (x, y) \ A x y" by(auto simp add: eventually_map_filter_on elim!: eventually_mono) from this[folded eq] obtain Q'' where Q'': "eventually Q'' G" and Q''P: "{y. \z. Q'' (y, z)} \ {y. \x. ?P (x, y)}" using G by(fastforce simp add: eventually_map_filter_on) have "eventually (inf Q'' ?Q) G" using Q'' Q' by(auto intro: eventually_conj simp add: inf_fun_def) then have "eventually Q' G" using Q''P by(auto elim!: eventually_mono simp add: Q'_def) moreover from Q' G have "\\<^sub>F y in ?Y2. \z. Q (y, z) \ B y z" by(auto simp add: eventually_map_filter_on elim!: eventually_mono) from this[unfolded eq] obtain P'' where P'': "eventually P'' F" and P''Q: "{y. \x. P'' (x, y)} \ {y. \z. ?Q (y, z)}" using F by(fastforce simp add: eventually_map_filter_on) have "eventually (inf P'' ?P) F" using P'' P' by(auto intro: eventually_conj simp add: inf_fun_def) then have "eventually P' F" using P''Q by(auto elim!: eventually_mono simp add: P'_def) ultimately show ?thesis by blast qed show "rel_filter (A OO B) ?X ?Z" proof let ?Y = "\Y. \X Z. eventually X ?X \ eventually Z ?Z \ (\(x, z). X x \ Z z \ (A OO B) x z) \ Y" have Y: "is_filter ?Y" proof show "?Y (\_. True)" by(auto simp add: le_fun_def intro: eventually_True) show "?Y (\x. P x \ Q x)" if "?Y P" "?Y Q" for P Q using that apply clarify apply(intro exI conjI; (elim eventually_rev_mp; fold imp_conjL; intro always_eventually allI; rule imp_refl)?) apply auto done show "?Y Q" if "?Y P" "\x. P x \ Q x" for P Q using that by blast qed define Y where "Y = Abs_filter ?Y" have eventually_Y: "eventually P Y \ ?Y P" for P using eventually_Abs_filter[OF Y, of P] by(simp add: Y_def) show YY: "\\<^sub>F (x, y) in Y. (A OO B) x y" using F G by(auto simp add: eventually_Y eventually_map_filter_on eventually_conj_iff intro!: eventually_True) have "?Y (\(x, z). P x \ (A OO B) x z) \ (\\<^sub>F (x, y) in F. P x \ A x y)" (is "?lhs = ?rhs") for P proof show ?lhs if ?rhs using G F that by(auto 4 3 intro: exI[where x="\_. True"] simp add: eventually_map_filter_on split_def) assume ?lhs then obtain X Z where "\\<^sub>F (x, y) in F. X x \ A x y" and "\\<^sub>F (x, y) in G. Z y \ B x y" and "(\(x, z). X x \ Z z \ (A OO B) x z) \ (\(x, z). P x \ (A OO B) x z)" using F G by(auto simp add: eventually_map_filter_on split_def) from step[OF this(1, 2)] this(3) show ?rhs by(clarsimp elim!: eventually_rev_mp simp add: le_fun_def)(fastforce intro: always_eventually) qed then show "map_filter_on ?AB fst Y = ?X" by(simp add: filter_eq_iff YY eventually_map_filter_on)(simp add: eventually_Y eventually_map_filter_on F G; simp add: split_def) have "?Y (\(x, z). P z \ (A OO B) x z) \ (\\<^sub>F (x, y) in G. P y \ B x y)" (is "?lhs = ?rhs") for P proof show ?lhs if ?rhs using G F that by(auto 4 3 intro: exI[where x="\_. True"] simp add: eventually_map_filter_on split_def) assume ?lhs then obtain X Z where "\\<^sub>F (x, y) in F. X x \ A x y" and "\\<^sub>F (x, y) in G. Z y \ B x y" and "(\(x, z). X x \ Z z \ (A OO B) x z) \ (\(x, z). P z \ (A OO B) x z)" using F G by(auto simp add: eventually_map_filter_on split_def) from step[OF this(1, 2)] this(3) show ?rhs by(clarsimp elim!: eventually_rev_mp simp add: le_fun_def)(fastforce intro: always_eventually) qed then show "map_filter_on ?AB snd Y = ?Z" by(simp add: filter_eq_iff YY eventually_map_filter_on)(simp add: eventually_Y eventually_map_filter_on F G; simp add: split_def) qed qed lemma filtermap_parametric: "((A ===> B) ===> rel_filter A ===> rel_filter B) filtermap filtermap" proof(intro rel_funI; erule rel_filter.cases; hypsubst) fix f g Z assume fg: "(A ===> B) f g" and Z: "\\<^sub>F (x, y) in Z. A x y" have "rel_filter B (map_filter_on {(x, y). A x y} (f \ fst) Z) (map_filter_on {(x, y). A x y} (g \ snd) Z)" (is "rel_filter _ ?F ?G") proof let ?Z = "map_filter_on {(x, y). A x y} (map_prod f g) Z" show "\\<^sub>F (x, y) in ?Z. B x y" using fg Z by(auto simp add: eventually_map_filter_on split_def elim!: eventually_mono rel_funD) have [simp]: "map_prod f g ` {p. A (fst p) (snd p)} \ {p. B (fst p) (snd p)}" using fg by(auto dest: rel_funD) show "map_filter_on {(x, y). B x y} fst ?Z = ?F" "map_filter_on {(x, y). B x y} snd ?Z = ?G" using Z by(auto simp add: map_filter_on_comp split_def) qed thus "rel_filter B (filtermap f (map_filter_on {(x, y). A x y} fst Z)) (filtermap g (map_filter_on {(x, y). A x y} snd Z))" using Z by(simp add: map_filter_on_UNIV[symmetric] map_filter_on_comp) qed lemma rel_filter_Grp: "rel_filter (Grp UNIV f) = Grp UNIV (filtermap f)" proof((intro antisym predicate2I; (elim GrpE; hypsubst)?), rule GrpI[OF _ UNIV_I]) fix F G assume "rel_filter (Grp UNIV f) F G" hence "rel_filter (=) (filtermap f F) (filtermap id G)" by(rule filtermap_parametric[THEN rel_funD, THEN rel_funD, rotated])(simp add: Grp_def rel_fun_def) thus "filtermap f F = G" by(simp add: rel_filter_eq) next fix F :: "'a filter" have "rel_filter (=) F F" by(simp add: rel_filter_eq) hence "rel_filter (Grp UNIV f) (filtermap id F) (filtermap f F)" by(rule filtermap_parametric[THEN rel_funD, THEN rel_funD, rotated])(simp add: Grp_def rel_fun_def) thus "rel_filter (Grp UNIV f) F (filtermap f F)" by simp qed lemma Quotient_filter [quot_map]: "Quotient R Abs Rep T \ Quotient (rel_filter R) (filtermap Abs) (filtermap Rep) (rel_filter T)" unfolding Quotient_alt_def5 rel_filter_eq[symmetric] rel_filter_Grp[symmetric] by(simp add: rel_filter_distr[symmetric] rel_filter_conversep[symmetric] rel_filter_mono) lemma left_total_rel_filter [transfer_rule]: "left_total A \ left_total (rel_filter A)" unfolding left_total_alt_def rel_filter_eq[symmetric] rel_filter_conversep[symmetric] rel_filter_distr by(rule rel_filter_mono) lemma right_total_rel_filter [transfer_rule]: "right_total A \ right_total (rel_filter A)" using left_total_rel_filter[of "A\\"] by(simp add: rel_filter_conversep) lemma bi_total_rel_filter [transfer_rule]: "bi_total A \ bi_total (rel_filter A)" unfolding bi_total_alt_def by(simp add: left_total_rel_filter right_total_rel_filter) lemma left_unique_rel_filter [transfer_rule]: "left_unique A \ left_unique (rel_filter A)" unfolding left_unique_alt_def rel_filter_eq[symmetric] rel_filter_conversep[symmetric] rel_filter_distr by(rule rel_filter_mono) lemma right_unique_rel_filter [transfer_rule]: "right_unique A \ right_unique (rel_filter A)" using left_unique_rel_filter[of "A\\"] by(simp add: rel_filter_conversep) lemma bi_unique_rel_filter [transfer_rule]: "bi_unique A \ bi_unique (rel_filter A)" by(simp add: bi_unique_alt_def left_unique_rel_filter right_unique_rel_filter) lemma eventually_parametric [transfer_rule]: "((A ===> (=)) ===> rel_filter A ===> (=)) eventually eventually" by(auto 4 4 intro!: rel_funI elim!: rel_filter.cases simp add: eventually_map_filter_on dest: rel_funD intro: always_eventually elim!: eventually_rev_mp) lemma frequently_parametric [transfer_rule]: "((A ===> (=)) ===> rel_filter A ===> (=)) frequently frequently" unfolding frequently_def[abs_def] by transfer_prover lemma is_filter_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" assumes [transfer_rule]: "bi_unique A" shows "(((A ===> (=)) ===> (=)) ===> (=)) is_filter is_filter" unfolding is_filter_def by transfer_prover lemma top_filter_parametric [transfer_rule]: "rel_filter A top top" if "bi_total A" proof let ?Z = "principal {(x, y). A x y}" show "\\<^sub>F (x, y) in ?Z. A x y" by(simp add: eventually_principal) show "map_filter_on {(x, y). A x y} fst ?Z = top" "map_filter_on {(x, y). A x y} snd ?Z = top" using that by(auto simp add: filter_eq_iff eventually_map_filter_on eventually_principal bi_total_def) qed lemma bot_filter_parametric [transfer_rule]: "rel_filter A bot bot" proof show "\\<^sub>F (x, y) in bot. A x y" by simp show "map_filter_on {(x, y). A x y} fst bot = bot" "map_filter_on {(x, y). A x y} snd bot = bot" by(simp_all add: filter_eq_iff eventually_map_filter_on) qed lemma principal_parametric [transfer_rule]: "(rel_set A ===> rel_filter A) principal principal" proof(rule rel_funI rel_filter.intros)+ fix S S' assume *: "rel_set A S S'" define SS' where "SS' = S \ S' \ {(x, y). A x y}" have SS': "SS' \ {(x, y). A x y}" and [simp]: "S = fst ` SS'" "S' = snd ` SS'" using * by(auto 4 3 dest: rel_setD1 rel_setD2 intro: rev_image_eqI simp add: SS'_def) let ?Z = "principal SS'" show "\\<^sub>F (x, y) in ?Z. A x y" using SS' by(auto simp add: eventually_principal) then show "map_filter_on {(x, y). A x y} fst ?Z = principal S" and "map_filter_on {(x, y). A x y} snd ?Z = principal S'" by(auto simp add: filter_eq_iff eventually_map_filter_on eventually_principal) qed lemma sup_filter_parametric [transfer_rule]: "(rel_filter A ===> rel_filter A ===> rel_filter A) sup sup" proof(intro rel_funI; elim rel_filter.cases; hypsubst) show "rel_filter A (map_filter_on {(x, y). A x y} fst FG \ map_filter_on {(x, y). A x y} fst FG') (map_filter_on {(x, y). A x y} snd FG \ map_filter_on {(x, y). A x y} snd FG')" (is "rel_filter _ (sup ?F ?G) (sup ?F' ?G')") if "\\<^sub>F (x, y) in FG. A x y" "\\<^sub>F (x, y) in FG'. A x y" for FG FG' proof let ?Z = "sup FG FG'" show "\\<^sub>F (x, y) in ?Z. A x y" by(simp add: eventually_sup that) then show "map_filter_on {(x, y). A x y} fst ?Z = sup ?F ?G" and "map_filter_on {(x, y). A x y} snd ?Z = sup ?F' ?G'" by(simp_all add: filter_eq_iff eventually_map_filter_on eventually_sup) qed qed lemma Sup_filter_parametric [transfer_rule]: "(rel_set (rel_filter A) ===> rel_filter A) Sup Sup" proof(rule rel_funI) fix S S' define SS' where "SS' = S \ S' \ {(F, G). rel_filter A F G}" assume "rel_set (rel_filter A) S S'" then have SS': "SS' \ {(F, G). rel_filter A F G}" and [simp]: "S = fst ` SS'" "S' = snd ` SS'" by(auto 4 3 dest: rel_setD1 rel_setD2 intro: rev_image_eqI simp add: SS'_def) from SS' obtain Z where Z: "\F G. (F, G) \ SS' \ (\\<^sub>F (x, y) in Z F G. A x y) \ id F = map_filter_on {(x, y). A x y} fst (Z F G) \ id G = map_filter_on {(x, y). A x y} snd (Z F G)" unfolding rel_filter.simps by atomize_elim((rule choice allI)+; auto) have id: "eventually P F = eventually P (id F)" "eventually Q G = eventually Q (id G)" if "(F, G) \ SS'" for P Q F G by simp_all show "rel_filter A (Sup S) (Sup S')" proof let ?Z = "\(F, G)\SS'. Z F G" show *: "\\<^sub>F (x, y) in ?Z. A x y" using Z by(auto simp add: eventually_Sup) show "map_filter_on {(x, y). A x y} fst ?Z = Sup S" "map_filter_on {(x, y). A x y} snd ?Z = Sup S'" unfolding filter_eq_iff by(auto 4 4 simp add: id eventually_Sup eventually_map_filter_on *[simplified eventually_Sup] simp del: id_apply dest: Z) qed qed context fixes A :: "'a \ 'b \ bool" assumes [transfer_rule]: "bi_unique A" begin lemma le_filter_parametric [transfer_rule]: "(rel_filter A ===> rel_filter A ===> (=)) (\) (\)" unfolding le_filter_def[abs_def] by transfer_prover lemma less_filter_parametric [transfer_rule]: "(rel_filter A ===> rel_filter A ===> (=)) (<) (<)" unfolding less_filter_def[abs_def] by transfer_prover context assumes [transfer_rule]: "bi_total A" begin lemma Inf_filter_parametric [transfer_rule]: "(rel_set (rel_filter A) ===> rel_filter A) Inf Inf" unfolding Inf_filter_def[abs_def] by transfer_prover lemma inf_filter_parametric [transfer_rule]: "(rel_filter A ===> rel_filter A ===> rel_filter A) inf inf" proof(intro rel_funI)+ fix F F' G G' assume [transfer_rule]: "rel_filter A F F'" "rel_filter A G G'" have "rel_filter A (Inf {F, G}) (Inf {F', G'})" by transfer_prover thus "rel_filter A (inf F G) (inf F' G')" by simp qed end end end -lemma prod_filter_parametric [transfer_rule]: includes lifting_syntax shows +context + includes lifting_syntax +begin + +lemma prod_filter_parametric [transfer_rule]: "(rel_filter R ===> rel_filter S ===> rel_filter (rel_prod R S)) prod_filter prod_filter" proof(intro rel_funI; elim rel_filter.cases; hypsubst) fix F G assume F: "\\<^sub>F (x, y) in F. R x y" and G: "\\<^sub>F (x, y) in G. S x y" show "rel_filter (rel_prod R S) (map_filter_on {(x, y). R x y} fst F \\<^sub>F map_filter_on {(x, y). S x y} fst G) (map_filter_on {(x, y). R x y} snd F \\<^sub>F map_filter_on {(x, y). S x y} snd G)" (is "rel_filter ?RS ?F ?G") proof let ?Z = "filtermap (\((a, b), (a', b')). ((a, a'), (b, b'))) (prod_filter F G)" show *: "\\<^sub>F (x, y) in ?Z. rel_prod R S x y" using F G by(auto simp add: eventually_filtermap split_beta eventually_prod_filter) show "map_filter_on {(x, y). ?RS x y} fst ?Z = ?F" using F G apply(clarsimp simp add: filter_eq_iff eventually_map_filter_on *) apply(simp add: eventually_filtermap split_beta eventually_prod_filter) apply(subst eventually_map_filter_on; simp)+ apply(rule iffI; clarsimp) subgoal for P P' P'' apply(rule exI[where x="\a. \b. P' (a, b) \ R a b"]; rule conjI) subgoal by(fastforce elim: eventually_rev_mp eventually_mono) subgoal by(rule exI[where x="\a. \b. P'' (a, b) \ S a b"])(fastforce elim: eventually_rev_mp eventually_mono) done subgoal by fastforce done show "map_filter_on {(x, y). ?RS x y} snd ?Z = ?G" using F G apply(clarsimp simp add: filter_eq_iff eventually_map_filter_on *) apply(simp add: eventually_filtermap split_beta eventually_prod_filter) apply(subst eventually_map_filter_on; simp)+ apply(rule iffI; clarsimp) subgoal for P P' P'' apply(rule exI[where x="\b. \a. P' (a, b) \ R a b"]; rule conjI) subgoal by(fastforce elim: eventually_rev_mp eventually_mono) subgoal by(rule exI[where x="\b. \a. P'' (a, b) \ S a b"])(fastforce elim: eventually_rev_mp eventually_mono) done subgoal by fastforce done qed qed +end + + text \Code generation for filters\ definition abstract_filter :: "(unit \ 'a filter) \ 'a filter" where [simp]: "abstract_filter f = f ()" code_datatype principal abstract_filter hide_const (open) abstract_filter declare [[code drop: filterlim prod_filter filtermap eventually "inf :: _ filter \ _" "sup :: _ filter \ _" "less_eq :: _ filter \ _" Abs_filter]] declare filterlim_principal [code] declare principal_prod_principal [code] declare filtermap_principal [code] declare filtercomap_principal [code] declare eventually_principal [code] declare inf_principal [code] declare sup_principal [code] declare principal_le_iff [code] lemma Rep_filter_iff_eventually [simp, code]: "Rep_filter F P \ eventually P F" by (simp add: eventually_def) lemma bot_eq_principal_empty [code]: "bot = principal {}" by simp lemma top_eq_principal_UNIV [code]: "top = principal UNIV" by simp instantiation filter :: (equal) equal begin definition equal_filter :: "'a filter \ 'a filter \ bool" where "equal_filter F F' \ F = F'" lemma equal_filter [code]: "HOL.equal (principal A) (principal B) \ A = B" by (simp add: equal_filter_def) instance by standard (simp add: equal_filter_def) end end diff --git a/src/HOL/Groups_List.thy b/src/HOL/Groups_List.thy --- a/src/HOL/Groups_List.thy +++ b/src/HOL/Groups_List.thy @@ -1,428 +1,432 @@ (* Author: Tobias Nipkow, TU Muenchen *) section \Sum and product over lists\ theory Groups_List imports List begin locale monoid_list = monoid begin definition F :: "'a list \ 'a" where eq_foldr [code]: "F xs = foldr f xs \<^bold>1" lemma Nil [simp]: "F [] = \<^bold>1" by (simp add: eq_foldr) lemma Cons [simp]: "F (x # xs) = x \<^bold>* F xs" by (simp add: eq_foldr) lemma append [simp]: "F (xs @ ys) = F xs \<^bold>* F ys" by (induct xs) (simp_all add: assoc) end locale comm_monoid_list = comm_monoid + monoid_list begin lemma rev [simp]: "F (rev xs) = F xs" by (simp add: eq_foldr foldr_fold fold_rev fun_eq_iff assoc left_commute) end locale comm_monoid_list_set = list: comm_monoid_list + set: comm_monoid_set begin lemma distinct_set_conv_list: "distinct xs \ set.F g (set xs) = list.F (map g xs)" by (induct xs) simp_all lemma set_conv_list [code]: "set.F g (set xs) = list.F (map g (remdups xs))" by (simp add: distinct_set_conv_list [symmetric]) end subsection \List summation\ context monoid_add begin sublocale sum_list: monoid_list plus 0 defines sum_list = sum_list.F .. end context comm_monoid_add begin sublocale sum_list: comm_monoid_list plus 0 rewrites "monoid_list.F plus 0 = sum_list" proof - show "comm_monoid_list plus 0" .. then interpret sum_list: comm_monoid_list plus 0 . from sum_list_def show "monoid_list.F plus 0 = sum_list" by simp qed sublocale sum: comm_monoid_list_set plus 0 rewrites "monoid_list.F plus 0 = sum_list" and "comm_monoid_set.F plus 0 = sum" proof - show "comm_monoid_list_set plus 0" .. then interpret sum: comm_monoid_list_set plus 0 . from sum_list_def show "monoid_list.F plus 0 = sum_list" by simp from sum_def show "comm_monoid_set.F plus 0 = sum" by (auto intro: sym) qed end text \Some syntactic sugar for summing a function over a list:\ syntax (ASCII) "_sum_list" :: "pttrn => 'a list => 'b => 'b" ("(3SUM _<-_. _)" [0, 51, 10] 10) syntax "_sum_list" :: "pttrn => 'a list => 'b => 'b" ("(3\_\_. _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\x\xs. b" == "CONST sum_list (CONST map (\x. b) xs)" text \TODO duplicates\ lemmas sum_list_simps = sum_list.Nil sum_list.Cons lemmas sum_list_append = sum_list.append lemmas sum_list_rev = sum_list.rev lemma (in monoid_add) fold_plus_sum_list_rev: "fold plus xs = plus (sum_list (rev xs))" proof fix x have "fold plus xs x = sum_list (rev xs @ [x])" by (simp add: foldr_conv_fold sum_list.eq_foldr) also have "\ = sum_list (rev xs) + x" by simp finally show "fold plus xs x = sum_list (rev xs) + x" . qed lemma (in comm_monoid_add) sum_list_map_remove1: "x \ set xs \ sum_list (map f xs) = f x + sum_list (map f (remove1 x xs))" by (induct xs) (auto simp add: ac_simps) lemma (in monoid_add) size_list_conv_sum_list: "size_list f xs = sum_list (map f xs) + size xs" by (induct xs) auto lemma (in monoid_add) length_concat: "length (concat xss) = sum_list (map length xss)" by (induct xss) simp_all lemma (in monoid_add) length_product_lists: "length (product_lists xss) = foldr (*) (map length xss) 1" proof (induct xss) case (Cons xs xss) then show ?case by (induct xs) (auto simp: length_concat o_def) qed simp lemma (in monoid_add) sum_list_map_filter: assumes "\x. x \ set xs \ \ P x \ f x = 0" shows "sum_list (map f (filter P xs)) = sum_list (map f xs)" using assms by (induct xs) auto lemma sum_list_filter_le_nat: fixes f :: "'a \ nat" shows "sum_list (map f (filter P xs)) \ sum_list (map f xs)" by(induction xs; simp) lemma (in comm_monoid_add) distinct_sum_list_conv_Sum: "distinct xs \ sum_list xs = Sum (set xs)" by (induct xs) simp_all lemma sum_list_upt[simp]: "m \ n \ sum_list [m.. {m..x. x \ set xs \ 0 \ x) \ 0 \ sum_list xs" by (induction xs) auto lemma sum_list_nonpos: "(\x. x \ set xs \ x \ 0) \ sum_list xs \ 0" by (induction xs) (auto simp: add_nonpos_nonpos) lemma sum_list_nonneg_eq_0_iff: "(\x. x \ set xs \ 0 \ x) \ sum_list xs = 0 \ (\x\ set xs. x = 0)" by (induction xs) (simp_all add: add_nonneg_eq_0_iff sum_list_nonneg) end context canonically_ordered_monoid_add begin lemma sum_list_eq_0_iff [simp]: "sum_list ns = 0 \ (\n \ set ns. n = 0)" by (simp add: sum_list_nonneg_eq_0_iff) lemma member_le_sum_list: "x \ set xs \ x \ sum_list xs" by (induction xs) (auto simp: add_increasing add_increasing2) lemma elem_le_sum_list: "k < size ns \ ns ! k \ sum_list (ns)" by (rule member_le_sum_list) simp end lemma (in ordered_cancel_comm_monoid_diff) sum_list_update: "k < size xs \ sum_list (xs[k := x]) = sum_list xs + x - xs ! k" apply(induction xs arbitrary:k) apply (auto simp: add_ac split: nat.split) apply(drule elem_le_sum_list) by (simp add: local.add_diff_assoc local.add_increasing) lemma (in monoid_add) sum_list_triv: "(\x\xs. r) = of_nat (length xs) * r" by (induct xs) (simp_all add: distrib_right) lemma (in monoid_add) sum_list_0 [simp]: "(\x\xs. 0) = 0" by (induct xs) (simp_all add: distrib_right) text\For non-Abelian groups \xs\ needs to be reversed on one side:\ lemma (in ab_group_add) uminus_sum_list_map: "- sum_list (map f xs) = sum_list (map (uminus \ f) xs)" by (induct xs) simp_all lemma (in comm_monoid_add) sum_list_addf: "(\x\xs. f x + g x) = sum_list (map f xs) + sum_list (map g xs)" by (induct xs) (simp_all add: algebra_simps) lemma (in ab_group_add) sum_list_subtractf: "(\x\xs. f x - g x) = sum_list (map f xs) - sum_list (map g xs)" by (induct xs) (simp_all add: algebra_simps) lemma (in semiring_0) sum_list_const_mult: "(\x\xs. c * f x) = c * (\x\xs. f x)" by (induct xs) (simp_all add: algebra_simps) lemma (in semiring_0) sum_list_mult_const: "(\x\xs. f x * c) = (\x\xs. f x) * c" by (induct xs) (simp_all add: algebra_simps) lemma (in ordered_ab_group_add_abs) sum_list_abs: "\sum_list xs\ \ sum_list (map abs xs)" by (induct xs) (simp_all add: order_trans [OF abs_triangle_ineq]) lemma sum_list_mono: fixes f g :: "'a \ 'b::{monoid_add, ordered_ab_semigroup_add}" shows "(\x. x \ set xs \ f x \ g x) \ (\x\xs. f x) \ (\x\xs. g x)" by (induct xs) (simp, simp add: add_mono) lemma sum_list_strict_mono: fixes f g :: "'a \ 'b::{monoid_add, strict_ordered_ab_semigroup_add}" shows "\ xs \ []; \x. x \ set xs \ f x < g x \ \ sum_list (map f xs) < sum_list (map g xs)" proof (induction xs) case Nil thus ?case by simp next case C: (Cons _ xs) show ?case proof (cases xs) case Nil thus ?thesis using C.prems by simp next case Cons thus ?thesis using C by(simp add: add_strict_mono) qed qed lemma (in monoid_add) sum_list_distinct_conv_sum_set: "distinct xs \ sum_list (map f xs) = sum f (set xs)" by (induct xs) simp_all lemma (in monoid_add) interv_sum_list_conv_sum_set_nat: "sum_list (map f [m..General equivalence between \<^const>\sum_list\ and \<^const>\sum\\ lemma (in monoid_add) sum_list_sum_nth: "sum_list xs = (\ i = 0 ..< length xs. xs ! i)" using interv_sum_list_conv_sum_set_nat [of "(!) xs" 0 "length xs"] by (simp add: map_nth) lemma sum_list_map_eq_sum_count: "sum_list (map f xs) = sum (\x. count_list xs x * f x) (set xs)" proof(induction xs) case (Cons x xs) show ?case (is "?l = ?r") proof cases assume "x \ set xs" have "?l = f x + (\x\set xs. count_list xs x * f x)" by (simp add: Cons.IH) also have "set xs = insert x (set xs - {x})" using \x \ set xs\by blast also have "f x + (\x\insert x (set xs - {x}). count_list xs x * f x) = ?r" by (simp add: sum.insert_remove eq_commute) finally show ?thesis . next assume "x \ set xs" hence "\xa. xa \ set xs \ x \ xa" by blast thus ?thesis by (simp add: Cons.IH \x \ set xs\) qed qed simp lemma sum_list_map_eq_sum_count2: assumes "set xs \ X" "finite X" shows "sum_list (map f xs) = sum (\x. count_list xs x * f x) X" proof- let ?F = "\x. count_list xs x * f x" have "sum ?F X = sum ?F (set xs \ (X - set xs))" using Un_absorb1[OF assms(1)] by(simp) also have "\ = sum ?F (set xs)" using assms(2) by(simp add: sum.union_disjoint[OF _ _ Diff_disjoint] del: Un_Diff_cancel) finally show ?thesis by(simp add:sum_list_map_eq_sum_count) qed lemma sum_list_nonneg: "(\x. x \ set xs \ (x :: 'a :: ordered_comm_monoid_add) \ 0) \ sum_list xs \ 0" by (induction xs) simp_all lemma sum_list_Suc: "sum_list (map (\x. Suc(f x)) xs) = sum_list (map f xs) + length xs" by(induction xs; simp) lemma (in monoid_add) sum_list_map_filter': "sum_list (map f (filter P xs)) = sum_list (map (\x. if P x then f x else 0) xs)" by (induction xs) simp_all text \Summation of a strictly ascending sequence with length \n\ can be upper-bounded by summation over \{0...\ lemma sorted_wrt_less_sum_mono_lowerbound: fixes f :: "nat \ ('b::ordered_comm_monoid_add)" assumes mono: "\x y. x\y \ f x \ f y" shows "sorted_wrt (<) ns \ (\i\{0.. (\i\ns. f i)" proof (induction ns rule: rev_induct) case Nil then show ?case by simp next case (snoc n ns) have "sum f {0.. sum_list (map f ns)" using snoc by (auto simp: sorted_wrt_append) also have "length ns \ n" using sorted_wrt_less_idx[OF snoc.prems(1), of "length ns"] by auto finally have "sum f {0.. sum_list (map f ns) + f n" using mono add_mono by blast thus ?case by simp qed subsection \Further facts about \<^const>\List.n_lists\\ lemma length_n_lists: "length (List.n_lists n xs) = length xs ^ n" by (induct n) (auto simp add: comp_def length_concat sum_list_triv) lemma distinct_n_lists: assumes "distinct xs" shows "distinct (List.n_lists n xs)" proof (rule card_distinct) from assms have card_length: "card (set xs) = length xs" by (rule distinct_card) have "card (set (List.n_lists n xs)) = card (set xs) ^ n" proof (induct n) case 0 then show ?case by simp next case (Suc n) moreover have "card (\ys\set (List.n_lists n xs). (\y. y # ys) ` set xs) = (\ys\set (List.n_lists n xs). card ((\y. y # ys) ` set xs))" by (rule card_UN_disjoint) auto moreover have "\ys. card ((\y. y # ys) ` set xs) = card (set xs)" by (rule card_image) (simp add: inj_on_def) ultimately show ?case by auto qed also have "\ = length xs ^ n" by (simp add: card_length) finally show "card (set (List.n_lists n xs)) = length (List.n_lists n xs)" by (simp add: length_n_lists) qed subsection \Tools setup\ lemmas sum_code = sum.set_conv_list lemma sum_set_upto_conv_sum_list_int [code_unfold]: "sum f (set [i..j::int]) = sum_list (map f [i..j])" by (simp add: interv_sum_list_conv_sum_set_int) lemma sum_set_upt_conv_sum_list_nat [code_unfold]: "sum f (set [m.. A ===> A) (+) (+)" - shows "(list_all2 A ===> A) sum_list sum_list" +begin + +lemma sum_list_transfer [transfer_rule]: + "(list_all2 A ===> A) sum_list sum_list" + if [transfer_rule]: "A 0 0" "(A ===> A ===> A) (+) (+)" unfolding sum_list.eq_foldr [abs_def] by transfer_prover +end + subsection \List product\ context monoid_mult begin sublocale prod_list: monoid_list times 1 defines prod_list = prod_list.F .. end context comm_monoid_mult begin sublocale prod_list: comm_monoid_list times 1 rewrites "monoid_list.F times 1 = prod_list" proof - show "comm_monoid_list times 1" .. then interpret prod_list: comm_monoid_list times 1 . from prod_list_def show "monoid_list.F times 1 = prod_list" by simp qed sublocale prod: comm_monoid_list_set times 1 rewrites "monoid_list.F times 1 = prod_list" and "comm_monoid_set.F times 1 = prod" proof - show "comm_monoid_list_set times 1" .. then interpret prod: comm_monoid_list_set times 1 . from prod_list_def show "monoid_list.F times 1 = prod_list" by simp from prod_def show "comm_monoid_set.F times 1 = prod" by (auto intro: sym) qed end lemma prod_list_zero_iff: "prod_list xs = 0 \ (0 :: 'a :: {semiring_no_zero_divisors, semiring_1}) \ set xs" by (induction xs) simp_all text \Some syntactic sugar:\ syntax (ASCII) "_prod_list" :: "pttrn => 'a list => 'b => 'b" ("(3PROD _<-_. _)" [0, 51, 10] 10) syntax "_prod_list" :: "pttrn => 'a list => 'b => 'b" ("(3\_\_. _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\x\xs. b" \ "CONST prod_list (CONST map (\x. b) xs)" end diff --git a/src/HOL/Int.thy b/src/HOL/Int.thy --- a/src/HOL/Int.thy +++ b/src/HOL/Int.thy @@ -1,1904 +1,1908 @@ (* Title: HOL/Int.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Tobias Nipkow, Florian Haftmann, TU Muenchen *) section \The Integers as Equivalence Classes over Pairs of Natural Numbers\ theory Int imports Equiv_Relations Power Quotient Fun_Def begin subsection \Definition of integers as a quotient type\ definition intrel :: "(nat \ nat) \ (nat \ nat) \ bool" where "intrel = (\(x, y) (u, v). x + v = u + y)" lemma intrel_iff [simp]: "intrel (x, y) (u, v) \ x + v = u + y" by (simp add: intrel_def) quotient_type int = "nat \ nat" / "intrel" morphisms Rep_Integ Abs_Integ proof (rule equivpI) show "reflp intrel" by (auto simp: reflp_def) show "symp intrel" by (auto simp: symp_def) show "transp intrel" by (auto simp: transp_def) qed lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]: "(\x y. z = Abs_Integ (x, y) \ P) \ P" by (induct z) auto subsection \Integers form a commutative ring\ instantiation int :: comm_ring_1 begin lift_definition zero_int :: "int" is "(0, 0)" . lift_definition one_int :: "int" is "(1, 0)" . lift_definition plus_int :: "int \ int \ int" is "\(x, y) (u, v). (x + u, y + v)" by clarsimp lift_definition uminus_int :: "int \ int" is "\(x, y). (y, x)" by clarsimp lift_definition minus_int :: "int \ int \ int" is "\(x, y) (u, v). (x + v, y + u)" by clarsimp lift_definition times_int :: "int \ int \ int" is "\(x, y) (u, v). (x*u + y*v, x*v + y*u)" proof (clarsimp) fix s t u v w x y z :: nat assume "s + v = u + t" and "w + z = y + x" then have "(s + v) * w + (u + t) * x + u * (w + z) + v * (y + x) = (u + t) * w + (s + v) * x + u * (y + x) + v * (w + z)" by simp then show "(s * w + t * x) + (u * z + v * y) = (u * y + v * z) + (s * x + t * w)" by (simp add: algebra_simps) qed instance by standard (transfer; clarsimp simp: algebra_simps)+ end abbreviation int :: "nat \ int" where "int \ of_nat" lemma int_def: "int n = Abs_Integ (n, 0)" by (induct n) (simp add: zero_int.abs_eq, simp add: one_int.abs_eq plus_int.abs_eq) -lemma int_transfer [transfer_rule]: "(rel_fun (=) pcr_int) (\n. (n, 0)) int" +lemma int_transfer [transfer_rule]: + includes lifting_syntax + shows "rel_fun (=) pcr_int (\n. (n, 0)) int" by (simp add: rel_fun_def int.pcr_cr_eq cr_int_def int_def) lemma int_diff_cases: obtains (diff) m n where "z = int m - int n" by transfer clarsimp subsection \Integers are totally ordered\ instantiation int :: linorder begin lift_definition less_eq_int :: "int \ int \ bool" is "\(x, y) (u, v). x + v \ u + y" by auto lift_definition less_int :: "int \ int \ bool" is "\(x, y) (u, v). x + v < u + y" by auto instance by standard (transfer, force)+ end instantiation int :: distrib_lattice begin definition "(inf :: int \ int \ int) = min" definition "(sup :: int \ int \ int) = max" instance by standard (auto simp add: inf_int_def sup_int_def max_min_distrib2) end subsection \Ordering properties of arithmetic operations\ instance int :: ordered_cancel_ab_semigroup_add proof fix i j k :: int show "i \ j \ k + i \ k + j" by transfer clarsimp qed text \Strict Monotonicity of Multiplication.\ text \Strict, in 1st argument; proof is by induction on \k > 0\.\ lemma zmult_zless_mono2_lemma: "i < j \ 0 < k \ int k * i < int k * j" for i j :: int proof (induct k) case 0 then show ?case by simp next case (Suc k) then show ?case by (cases "k = 0") (simp_all add: distrib_right add_strict_mono) qed lemma zero_le_imp_eq_int: "0 \ k \ \n. k = int n" for k :: int apply transfer apply clarsimp apply (rule_tac x="a - b" in exI) apply simp done lemma zero_less_imp_eq_int: "0 < k \ \n>0. k = int n" for k :: int apply transfer apply clarsimp apply (rule_tac x="a - b" in exI) apply simp done lemma zmult_zless_mono2: "i < j \ 0 < k \ k * i < k * j" for i j k :: int by (drule zero_less_imp_eq_int) (auto simp add: zmult_zless_mono2_lemma) text \The integers form an ordered integral domain.\ instantiation int :: linordered_idom begin definition zabs_def: "\i::int\ = (if i < 0 then - i else i)" definition zsgn_def: "sgn (i::int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)" instance proof fix i j k :: int show "i < j \ 0 < k \ k * i < k * j" by (rule zmult_zless_mono2) show "\i\ = (if i < 0 then -i else i)" by (simp only: zabs_def) show "sgn (i::int) = (if i=0 then 0 else if 0 w + 1 \ z" for w z :: int by transfer clarsimp lemma zless_iff_Suc_zadd: "w < z \ (\n. z = w + int (Suc n))" for w z :: int apply transfer apply auto apply (rename_tac a b c d) apply (rule_tac x="c+b - Suc(a+d)" in exI) apply arith done lemma zabs_less_one_iff [simp]: "\z\ < 1 \ z = 0" (is "?lhs \ ?rhs") for z :: int proof assume ?rhs then show ?lhs by simp next assume ?lhs with zless_imp_add1_zle [of "\z\" 1] have "\z\ + 1 \ 1" by simp then have "\z\ \ 0" by simp then show ?rhs by simp qed subsection \Embedding of the Integers into any \ring_1\: \of_int\\ context ring_1 begin lift_definition of_int :: "int \ 'a" is "\(i, j). of_nat i - of_nat j" by (clarsimp simp add: diff_eq_eq eq_diff_eq diff_add_eq of_nat_add [symmetric] simp del: of_nat_add) lemma of_int_0 [simp]: "of_int 0 = 0" by transfer simp lemma of_int_1 [simp]: "of_int 1 = 1" by transfer simp lemma of_int_add [simp]: "of_int (w + z) = of_int w + of_int z" by transfer (clarsimp simp add: algebra_simps) lemma of_int_minus [simp]: "of_int (- z) = - (of_int z)" by (transfer fixing: uminus) clarsimp lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z" using of_int_add [of w "- z"] by simp lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z" by (transfer fixing: times) (clarsimp simp add: algebra_simps) lemma mult_of_int_commute: "of_int x * y = y * of_int x" by (transfer fixing: times) (auto simp: algebra_simps mult_of_nat_commute) text \Collapse nested embeddings.\ lemma of_int_of_nat_eq [simp]: "of_int (int n) = of_nat n" by (induct n) auto lemma of_int_numeral [simp, code_post]: "of_int (numeral k) = numeral k" by (simp add: of_nat_numeral [symmetric] of_int_of_nat_eq [symmetric]) lemma of_int_neg_numeral [code_post]: "of_int (- numeral k) = - numeral k" by simp lemma of_int_power [simp]: "of_int (z ^ n) = of_int z ^ n" by (induct n) simp_all lemma of_int_of_bool [simp]: "of_int (of_bool P) = of_bool P" by auto end context ring_char_0 begin lemma of_int_eq_iff [simp]: "of_int w = of_int z \ w = z" by transfer (clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add) text \Special cases where either operand is zero.\ lemma of_int_eq_0_iff [simp]: "of_int z = 0 \ z = 0" using of_int_eq_iff [of z 0] by simp lemma of_int_0_eq_iff [simp]: "0 = of_int z \ z = 0" using of_int_eq_iff [of 0 z] by simp lemma of_int_eq_1_iff [iff]: "of_int z = 1 \ z = 1" using of_int_eq_iff [of z 1] by simp lemma numeral_power_eq_of_int_cancel_iff [simp]: "numeral x ^ n = of_int y \ numeral x ^ n = y" using of_int_eq_iff[of "numeral x ^ n" y, unfolded of_int_numeral of_int_power] . lemma of_int_eq_numeral_power_cancel_iff [simp]: "of_int y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_of_int_cancel_iff [of x n y] by (metis (mono_tags)) lemma neg_numeral_power_eq_of_int_cancel_iff [simp]: "(- numeral x) ^ n = of_int y \ (- numeral x) ^ n = y" using of_int_eq_iff[of "(- numeral x) ^ n" y] by simp lemma of_int_eq_neg_numeral_power_cancel_iff [simp]: "of_int y = (- numeral x) ^ n \ y = (- numeral x) ^ n" using neg_numeral_power_eq_of_int_cancel_iff[of x n y] by (metis (mono_tags)) lemma of_int_eq_of_int_power_cancel_iff[simp]: "(of_int b) ^ w = of_int x \ b ^ w = x" by (metis of_int_power of_int_eq_iff) lemma of_int_power_eq_of_int_cancel_iff[simp]: "of_int x = (of_int b) ^ w \ x = b ^ w" by (metis of_int_eq_of_int_power_cancel_iff) end context linordered_idom begin text \Every \linordered_idom\ has characteristic zero.\ subclass ring_char_0 .. lemma of_int_le_iff [simp]: "of_int w \ of_int z \ w \ z" by (transfer fixing: less_eq) (clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add) lemma of_int_less_iff [simp]: "of_int w < of_int z \ w < z" by (simp add: less_le order_less_le) lemma of_int_0_le_iff [simp]: "0 \ of_int z \ 0 \ z" using of_int_le_iff [of 0 z] by simp lemma of_int_le_0_iff [simp]: "of_int z \ 0 \ z \ 0" using of_int_le_iff [of z 0] by simp lemma of_int_0_less_iff [simp]: "0 < of_int z \ 0 < z" using of_int_less_iff [of 0 z] by simp lemma of_int_less_0_iff [simp]: "of_int z < 0 \ z < 0" using of_int_less_iff [of z 0] by simp lemma of_int_1_le_iff [simp]: "1 \ of_int z \ 1 \ z" using of_int_le_iff [of 1 z] by simp lemma of_int_le_1_iff [simp]: "of_int z \ 1 \ z \ 1" using of_int_le_iff [of z 1] by simp lemma of_int_1_less_iff [simp]: "1 < of_int z \ 1 < z" using of_int_less_iff [of 1 z] by simp lemma of_int_less_1_iff [simp]: "of_int z < 1 \ z < 1" using of_int_less_iff [of z 1] by simp lemma of_int_pos: "z > 0 \ of_int z > 0" by simp lemma of_int_nonneg: "z \ 0 \ of_int z \ 0" by simp lemma of_int_abs [simp]: "of_int \x\ = \of_int x\" by (auto simp add: abs_if) lemma of_int_lessD: assumes "\of_int n\ < x" shows "n = 0 \ x > 1" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "\n\ \ 0" by simp then have "\n\ > 0" by simp then have "\n\ \ 1" using zless_imp_add1_zle [of 0 "\n\"] by simp then have "\of_int n\ \ 1" unfolding of_int_1_le_iff [of "\n\", symmetric] by simp then have "1 < x" using assms by (rule le_less_trans) then show ?thesis .. qed lemma of_int_leD: assumes "\of_int n\ \ x" shows "n = 0 \ 1 \ x" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "\n\ \ 0" by simp then have "\n\ > 0" by simp then have "\n\ \ 1" using zless_imp_add1_zle [of 0 "\n\"] by simp then have "\of_int n\ \ 1" unfolding of_int_1_le_iff [of "\n\", symmetric] by simp then have "1 \ x" using assms by (rule order_trans) then show ?thesis .. qed lemma numeral_power_le_of_int_cancel_iff [simp]: "numeral x ^ n \ of_int a \ numeral x ^ n \ a" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_le_iff) lemma of_int_le_numeral_power_cancel_iff [simp]: "of_int a \ numeral x ^ n \ a \ numeral x ^ n" by (metis (mono_tags) local.numeral_power_eq_of_int_cancel_iff of_int_le_iff) lemma numeral_power_less_of_int_cancel_iff [simp]: "numeral x ^ n < of_int a \ numeral x ^ n < a" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_less_iff) lemma of_int_less_numeral_power_cancel_iff [simp]: "of_int a < numeral x ^ n \ a < numeral x ^ n" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_less_iff) lemma neg_numeral_power_le_of_int_cancel_iff [simp]: "(- numeral x) ^ n \ of_int a \ (- numeral x) ^ n \ a" by (metis (mono_tags) of_int_le_iff of_int_neg_numeral of_int_power) lemma of_int_le_neg_numeral_power_cancel_iff [simp]: "of_int a \ (- numeral x) ^ n \ a \ (- numeral x) ^ n" by (metis (mono_tags) of_int_le_iff of_int_neg_numeral of_int_power) lemma neg_numeral_power_less_of_int_cancel_iff [simp]: "(- numeral x) ^ n < of_int a \ (- numeral x) ^ n < a" using of_int_less_iff[of "(- numeral x) ^ n" a] by simp lemma of_int_less_neg_numeral_power_cancel_iff [simp]: "of_int a < (- numeral x) ^ n \ a < (- numeral x::int) ^ n" using of_int_less_iff[of a "(- numeral x) ^ n"] by simp lemma of_int_le_of_int_power_cancel_iff[simp]: "(of_int b) ^ w \ of_int x \ b ^ w \ x" by (metis (mono_tags) of_int_le_iff of_int_power) lemma of_int_power_le_of_int_cancel_iff[simp]: "of_int x \ (of_int b) ^ w\ x \ b ^ w" by (metis (mono_tags) of_int_le_iff of_int_power) lemma of_int_less_of_int_power_cancel_iff[simp]: "(of_int b) ^ w < of_int x \ b ^ w < x" by (metis (mono_tags) of_int_less_iff of_int_power) lemma of_int_power_less_of_int_cancel_iff[simp]: "of_int x < (of_int b) ^ w\ x < b ^ w" by (metis (mono_tags) of_int_less_iff of_int_power) lemma of_int_max: "of_int (max x y) = max (of_int x) (of_int y)" by (auto simp: max_def) lemma of_int_min: "of_int (min x y) = min (of_int x) (of_int y)" by (auto simp: min_def) end context division_ring begin lemmas mult_inverse_of_int_commute = mult_commute_imp_mult_inverse_commute[OF mult_of_int_commute] end text \Comparisons involving \<^term>\of_int\.\ lemma of_int_eq_numeral_iff [iff]: "of_int z = (numeral n :: 'a::ring_char_0) \ z = numeral n" using of_int_eq_iff by fastforce lemma of_int_le_numeral_iff [simp]: "of_int z \ (numeral n :: 'a::linordered_idom) \ z \ numeral n" using of_int_le_iff [of z "numeral n"] by simp lemma of_int_numeral_le_iff [simp]: "(numeral n :: 'a::linordered_idom) \ of_int z \ numeral n \ z" using of_int_le_iff [of "numeral n"] by simp lemma of_int_less_numeral_iff [simp]: "of_int z < (numeral n :: 'a::linordered_idom) \ z < numeral n" using of_int_less_iff [of z "numeral n"] by simp lemma of_int_numeral_less_iff [simp]: "(numeral n :: 'a::linordered_idom) < of_int z \ numeral n < z" using of_int_less_iff [of "numeral n" z] by simp lemma of_nat_less_of_int_iff: "(of_nat n::'a::linordered_idom) < of_int x \ int n < x" by (metis of_int_of_nat_eq of_int_less_iff) lemma of_int_eq_id [simp]: "of_int = id" proof show "of_int z = id z" for z by (cases z rule: int_diff_cases) simp qed instance int :: no_top apply standard apply (rule_tac x="x + 1" in exI) apply simp done instance int :: no_bot apply standard apply (rule_tac x="x - 1" in exI) apply simp done subsection \Magnitude of an Integer, as a Natural Number: \nat\\ lift_definition nat :: "int \ nat" is "\(x, y). x - y" by auto lemma nat_int [simp]: "nat (int n) = n" by transfer simp lemma int_nat_eq [simp]: "int (nat z) = (if 0 \ z then z else 0)" by transfer clarsimp lemma nat_0_le: "0 \ z \ int (nat z) = z" by simp lemma nat_le_0 [simp]: "z \ 0 \ nat z = 0" by transfer clarsimp lemma nat_le_eq_zle: "0 < w \ 0 \ z \ nat w \ nat z \ w \ z" by transfer (clarsimp, arith) text \An alternative condition is \<^term>\0 \ w\.\ lemma nat_mono_iff: "0 < z \ nat w < nat z \ w < z" by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) lemma nat_less_eq_zless: "0 \ w \ nat w < nat z \ w < z" by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) lemma zless_nat_conj [simp]: "nat w < nat z \ 0 < z \ w < z" by transfer (clarsimp, arith) lemma nonneg_int_cases: assumes "0 \ k" obtains n where "k = int n" proof - from assms have "k = int (nat k)" by simp then show thesis by (rule that) qed lemma pos_int_cases: assumes "0 < k" obtains n where "k = int n" and "n > 0" proof - from assms have "0 \ k" by simp then obtain n where "k = int n" by (rule nonneg_int_cases) moreover have "n > 0" using \k = int n\ assms by simp ultimately show thesis by (rule that) qed lemma nonpos_int_cases: assumes "k \ 0" obtains n where "k = - int n" proof - from assms have "- k \ 0" by simp then obtain n where "- k = int n" by (rule nonneg_int_cases) then have "k = - int n" by simp then show thesis by (rule that) qed lemma neg_int_cases: assumes "k < 0" obtains n where "k = - int n" and "n > 0" proof - from assms have "- k > 0" by simp then obtain n where "- k = int n" and "- k > 0" by (blast elim: pos_int_cases) then have "k = - int n" and "n > 0" by simp_all then show thesis by (rule that) qed lemma nat_eq_iff: "nat w = m \ (if 0 \ w then w = int m else m = 0)" by transfer (clarsimp simp add: le_imp_diff_is_add) lemma nat_eq_iff2: "m = nat w \ (if 0 \ w then w = int m else m = 0)" using nat_eq_iff [of w m] by auto lemma nat_0 [simp]: "nat 0 = 0" by (simp add: nat_eq_iff) lemma nat_1 [simp]: "nat 1 = Suc 0" by (simp add: nat_eq_iff) lemma nat_numeral [simp]: "nat (numeral k) = numeral k" by (simp add: nat_eq_iff) lemma nat_neg_numeral [simp]: "nat (- numeral k) = 0" by simp lemma nat_2: "nat 2 = Suc (Suc 0)" by simp lemma nat_less_iff: "0 \ w \ nat w < m \ w < of_nat m" by transfer (clarsimp, arith) lemma nat_le_iff: "nat x \ n \ x \ int n" by transfer (clarsimp simp add: le_diff_conv) lemma nat_mono: "x \ y \ nat x \ nat y" by transfer auto lemma nat_0_iff[simp]: "nat i = 0 \ i \ 0" for i :: int by transfer clarsimp lemma int_eq_iff: "of_nat m = z \ m = nat z \ 0 \ z" by (auto simp add: nat_eq_iff2) lemma zero_less_nat_eq [simp]: "0 < nat z \ 0 < z" using zless_nat_conj [of 0] by auto lemma nat_add_distrib: "0 \ z \ 0 \ z' \ nat (z + z') = nat z + nat z'" by transfer clarsimp lemma nat_diff_distrib': "0 \ x \ 0 \ y \ nat (x - y) = nat x - nat y" by transfer clarsimp lemma nat_diff_distrib: "0 \ z' \ z' \ z \ nat (z - z') = nat z - nat z'" by (rule nat_diff_distrib') auto lemma nat_zminus_int [simp]: "nat (- int n) = 0" by transfer simp lemma le_nat_iff: "k \ 0 \ n \ nat k \ int n \ k" by transfer auto lemma zless_nat_eq_int_zless: "m < nat z \ int m < z" by transfer (clarsimp simp add: less_diff_conv) lemma (in ring_1) of_nat_nat [simp]: "0 \ z \ of_nat (nat z) = of_int z" by transfer (clarsimp simp add: of_nat_diff) lemma diff_nat_numeral [simp]: "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')" by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral) lemma nat_abs_triangle_ineq: "nat \k + l\ \ nat \k\ + nat \l\" by (simp add: nat_add_distrib [symmetric] nat_le_eq_zle abs_triangle_ineq) lemma nat_of_bool [simp]: "nat (of_bool P) = of_bool P" by auto lemma split_nat [arith_split]: "P (nat i) \ ((\n. i = int n \ P n) \ (i < 0 \ P 0))" (is "?P = (?L \ ?R)") for i :: int proof (cases "i < 0") case True then show ?thesis by auto next case False have "?P = ?L" proof assume ?P then show ?L using False by auto next assume ?L moreover from False have "int (nat i) = i" by (simp add: not_less) ultimately show ?P by simp qed with False show ?thesis by simp qed lemma all_nat: "(\x. P x) \ (\x\0. P (nat x))" by (auto split: split_nat) lemma ex_nat: "(\x. P x) \ (\x. 0 \ x \ P (nat x))" proof assume "\x. P x" then obtain x where "P x" .. then have "int x \ 0 \ P (nat (int x))" by simp then show "\x\0. P (nat x)" .. next assume "\x\0. P (nat x)" then show "\x. P x" by auto qed text \For termination proofs:\ lemma measure_function_int[measure_function]: "is_measure (nat \ abs)" .. subsection \Lemmas about the Function \<^term>\of_nat\ and Orderings\ lemma negative_zless_0: "- (int (Suc n)) < (0 :: int)" by (simp add: order_less_le del: of_nat_Suc) lemma negative_zless [iff]: "- (int (Suc n)) < int m" by (rule negative_zless_0 [THEN order_less_le_trans], simp) lemma negative_zle_0: "- int n \ 0" by (simp add: minus_le_iff) lemma negative_zle [iff]: "- int n \ int m" by (rule order_trans [OF negative_zle_0 of_nat_0_le_iff]) lemma not_zle_0_negative [simp]: "\ 0 \ - int (Suc n)" by (subst le_minus_iff) (simp del: of_nat_Suc) lemma int_zle_neg: "int n \ - int m \ n = 0 \ m = 0" by transfer simp lemma not_int_zless_negative [simp]: "\ int n < - int m" by (simp add: linorder_not_less) lemma negative_eq_positive [simp]: "- int n = of_nat m \ n = 0 \ m = 0" by (force simp add: order_eq_iff [of "- of_nat n"] int_zle_neg) lemma zle_iff_zadd: "w \ z \ (\n. z = w + int n)" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by auto next assume ?lhs then have "0 \ z - w" by simp then obtain n where "z - w = int n" using zero_le_imp_eq_int [of "z - w"] by blast then have "z = w + int n" by simp then show ?rhs .. qed lemma zadd_int_left: "int m + (int n + z) = int (m + n) + z" by simp text \ This version is proved for all ordered rings, not just integers! It is proved here because attribute \arith_split\ is not available in theory \Rings\. But is it really better than just rewriting with \abs_if\? \ lemma abs_split [arith_split, no_atp]: "P \a\ \ (0 \ a \ P a) \ (a < 0 \ P (- a))" for a :: "'a::linordered_idom" by (force dest: order_less_le_trans simp add: abs_if linorder_not_less) lemma negD: "x < 0 \ \n. x = - (int (Suc n))" apply transfer apply clarsimp apply (rule_tac x="b - Suc a" in exI) apply arith done subsection \Cases and induction\ text \ Now we replace the case analysis rule by a more conventional one: whether an integer is negative or not. \ text \This version is symmetric in the two subgoals.\ lemma int_cases2 [case_names nonneg nonpos, cases type: int]: "(\n. z = int n \ P) \ (\n. z = - (int n) \ P) \ P" by (cases "z < 0") (auto simp add: linorder_not_less dest!: negD nat_0_le [THEN sym]) text \This is the default, with a negative case.\ lemma int_cases [case_names nonneg neg, cases type: int]: "(\n. z = int n \ P) \ (\n. z = - (int (Suc n)) \ P) \ P" apply (cases "z < 0") apply (blast dest!: negD) apply (simp add: linorder_not_less del: of_nat_Suc) apply auto apply (blast dest: nat_0_le [THEN sym]) done lemma int_cases3 [case_names zero pos neg]: fixes k :: int assumes "k = 0 \ P" and "\n. k = int n \ n > 0 \ P" and "\n. k = - int n \ n > 0 \ P" shows "P" proof (cases k "0::int" rule: linorder_cases) case equal with assms(1) show P by simp next case greater then have *: "nat k > 0" by simp moreover from * have "k = int (nat k)" by auto ultimately show P using assms(2) by blast next case less then have *: "nat (- k) > 0" by simp moreover from * have "k = - int (nat (- k))" by auto ultimately show P using assms(3) by blast qed lemma int_of_nat_induct [case_names nonneg neg, induct type: int]: "(\n. P (int n)) \ (\n. P (- (int (Suc n)))) \ P z" by (cases z) auto lemma sgn_mult_dvd_iff [simp]: "sgn r * l dvd k \ l dvd k \ (r = 0 \ k = 0)" for k l r :: int by (cases r rule: int_cases3) auto lemma mult_sgn_dvd_iff [simp]: "l * sgn r dvd k \ l dvd k \ (r = 0 \ k = 0)" for k l r :: int using sgn_mult_dvd_iff [of r l k] by (simp add: ac_simps) lemma dvd_sgn_mult_iff [simp]: "l dvd sgn r * k \ l dvd k \ r = 0" for k l r :: int by (cases r rule: int_cases3) simp_all lemma dvd_mult_sgn_iff [simp]: "l dvd k * sgn r \ l dvd k \ r = 0" for k l r :: int using dvd_sgn_mult_iff [of l r k] by (simp add: ac_simps) lemma int_sgnE: fixes k :: int obtains n and l where "k = sgn l * int n" proof - have "k = sgn k * int (nat \k\)" by (simp add: sgn_mult_abs) then show ?thesis .. qed subsubsection \Binary comparisons\ text \Preliminaries\ lemma le_imp_0_less: fixes z :: int assumes le: "0 \ z" shows "0 < 1 + z" proof - have "0 \ z" by fact also have "\ < z + 1" by (rule less_add_one) also have "\ = 1 + z" by (simp add: ac_simps) finally show "0 < 1 + z" . qed lemma odd_less_0_iff: "1 + z + z < 0 \ z < 0" for z :: int proof (cases z) case (nonneg n) then show ?thesis by (simp add: linorder_not_less add.assoc add_increasing le_imp_0_less [THEN order_less_imp_le]) next case (neg n) then show ?thesis by (simp del: of_nat_Suc of_nat_add of_nat_1 add: algebra_simps of_nat_1 [where 'a=int, symmetric] of_nat_add [symmetric]) qed subsubsection \Comparisons, for Ordered Rings\ lemma odd_nonzero: "1 + z + z \ 0" for z :: int proof (cases z) case (nonneg n) have le: "0 \ z + z" by (simp add: nonneg add_increasing) then show ?thesis using le_imp_0_less [OF le] by (auto simp: ac_simps) next case (neg n) show ?thesis proof assume eq: "1 + z + z = 0" have "0 < 1 + (int n + int n)" by (simp add: le_imp_0_less add_increasing) also have "\ = - (1 + z + z)" by (simp add: neg add.assoc [symmetric]) also have "\ = 0" by (simp add: eq) finally have "0<0" .. then show False by blast qed qed subsection \The Set of Integers\ context ring_1 begin definition Ints :: "'a set" ("\") where "\ = range of_int" lemma Ints_of_int [simp]: "of_int z \ \" by (simp add: Ints_def) lemma Ints_of_nat [simp]: "of_nat n \ \" using Ints_of_int [of "of_nat n"] by simp lemma Ints_0 [simp]: "0 \ \" using Ints_of_int [of "0"] by simp lemma Ints_1 [simp]: "1 \ \" using Ints_of_int [of "1"] by simp lemma Ints_numeral [simp]: "numeral n \ \" by (subst of_nat_numeral [symmetric], rule Ints_of_nat) lemma Ints_add [simp]: "a \ \ \ b \ \ \ a + b \ \" apply (auto simp add: Ints_def) apply (rule range_eqI) apply (rule of_int_add [symmetric]) done lemma Ints_minus [simp]: "a \ \ \ -a \ \" apply (auto simp add: Ints_def) apply (rule range_eqI) apply (rule of_int_minus [symmetric]) done lemma minus_in_Ints_iff: "-x \ \ \ x \ \" using Ints_minus[of x] Ints_minus[of "-x"] by auto lemma Ints_diff [simp]: "a \ \ \ b \ \ \ a - b \ \" apply (auto simp add: Ints_def) apply (rule range_eqI) apply (rule of_int_diff [symmetric]) done lemma Ints_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" apply (auto simp add: Ints_def) apply (rule range_eqI) apply (rule of_int_mult [symmetric]) done lemma Ints_power [simp]: "a \ \ \ a ^ n \ \" by (induct n) simp_all lemma Ints_cases [cases set: Ints]: assumes "q \ \" obtains (of_int) z where "q = of_int z" unfolding Ints_def proof - from \q \ \\ have "q \ range of_int" unfolding Ints_def . then obtain z where "q = of_int z" .. then show thesis .. qed lemma Ints_induct [case_names of_int, induct set: Ints]: "q \ \ \ (\z. P (of_int z)) \ P q" by (rule Ints_cases) auto lemma Nats_subset_Ints: "\ \ \" unfolding Nats_def Ints_def by (rule subsetI, elim imageE, hypsubst, subst of_int_of_nat_eq[symmetric], rule imageI) simp_all lemma Nats_altdef1: "\ = {of_int n |n. n \ 0}" proof (intro subsetI equalityI) fix x :: 'a assume "x \ {of_int n |n. n \ 0}" then obtain n where "x = of_int n" "n \ 0" by (auto elim!: Ints_cases) then have "x = of_nat (nat n)" by (subst of_nat_nat) simp_all then show "x \ \" by simp next fix x :: 'a assume "x \ \" then obtain n where "x = of_nat n" by (auto elim!: Nats_cases) then have "x = of_int (int n)" by simp also have "int n \ 0" by simp then have "of_int (int n) \ {of_int n |n. n \ 0}" by blast finally show "x \ {of_int n |n. n \ 0}" . qed end lemma (in linordered_idom) Ints_abs [simp]: shows "a \ \ \ abs a \ \" by (auto simp: abs_if) lemma (in linordered_idom) Nats_altdef2: "\ = {n \ \. n \ 0}" proof (intro subsetI equalityI) fix x :: 'a assume "x \ {n \ \. n \ 0}" then obtain n where "x = of_int n" "n \ 0" by (auto elim!: Ints_cases) then have "x = of_nat (nat n)" by (subst of_nat_nat) simp_all then show "x \ \" by simp qed (auto elim!: Nats_cases) lemma (in idom_divide) of_int_divide_in_Ints: "of_int a div of_int b \ \" if "b dvd a" proof - from that obtain c where "a = b * c" .. then show ?thesis by (cases "of_int b = 0") simp_all qed text \The premise involving \<^term>\Ints\ prevents \<^term>\a = 1/2\.\ lemma Ints_double_eq_0_iff: fixes a :: "'a::ring_char_0" assumes in_Ints: "a \ \" shows "a + a = 0 \ a = 0" (is "?lhs \ ?rhs") proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. show ?thesis proof assume ?rhs then show ?lhs by simp next assume ?lhs with a have "of_int (z + z) = (of_int 0 :: 'a)" by simp then have "z + z = 0" by (simp only: of_int_eq_iff) then have "z = 0" by (simp only: double_zero) with a show ?rhs by simp qed qed lemma Ints_odd_nonzero: fixes a :: "'a::ring_char_0" assumes in_Ints: "a \ \" shows "1 + a + a \ 0" proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. show ?thesis proof assume "1 + a + a = 0" with a have "of_int (1 + z + z) = (of_int 0 :: 'a)" by simp then have "1 + z + z = 0" by (simp only: of_int_eq_iff) with odd_nonzero show False by blast qed qed lemma Nats_numeral [simp]: "numeral w \ \" using of_nat_in_Nats [of "numeral w"] by simp lemma Ints_odd_less_0: fixes a :: "'a::linordered_idom" assumes in_Ints: "a \ \" shows "1 + a + a < 0 \ a < 0" proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. with a have "1 + a + a < 0 \ of_int (1 + z + z) < (of_int 0 :: 'a)" by simp also have "\ \ z < 0" by (simp only: of_int_less_iff odd_less_0_iff) also have "\ \ a < 0" by (simp add: a) finally show ?thesis . qed subsection \\<^term>\sum\ and \<^term>\prod\\ context semiring_1 begin lemma of_nat_sum [simp]: "of_nat (sum f A) = (\x\A. of_nat (f x))" by (induction A rule: infinite_finite_induct) auto end context ring_1 begin lemma of_int_sum [simp]: "of_int (sum f A) = (\x\A. of_int (f x))" by (induction A rule: infinite_finite_induct) auto end context comm_semiring_1 begin lemma of_nat_prod [simp]: "of_nat (prod f A) = (\x\A. of_nat (f x))" by (induction A rule: infinite_finite_induct) auto end context comm_ring_1 begin lemma of_int_prod [simp]: "of_int (prod f A) = (\x\A. of_int (f x))" by (induction A rule: infinite_finite_induct) auto end subsection \Setting up simplification procedures\ ML_file \Tools/int_arith.ML\ declaration \K ( Lin_Arith.add_discrete_type \<^type_name>\Int.int\ #> Lin_Arith.add_lessD @{thm zless_imp_add1_zle} #> Lin_Arith.add_inj_thms @{thms of_nat_le_iff [THEN iffD2] of_nat_eq_iff [THEN iffD2]} #> Lin_Arith.add_inj_const (\<^const_name>\of_nat\, \<^typ>\nat \ int\) #> Lin_Arith.add_simps @{thms of_int_0 of_int_1 of_int_add of_int_mult of_int_numeral of_int_neg_numeral nat_0 nat_1 diff_nat_numeral nat_numeral neg_less_iff_less True_implies_equals distrib_left [where a = "numeral v" for v] distrib_left [where a = "- numeral v" for v] div_by_1 div_0 times_divide_eq_right times_divide_eq_left minus_divide_left [THEN sym] minus_divide_right [THEN sym] add_divide_distrib diff_divide_distrib of_int_minus of_int_diff of_int_of_nat_eq} #> Lin_Arith.add_simprocs [Int_Arith.zero_one_idom_simproc] )\ simproc_setup fast_arith ("(m::'a::linordered_idom) < n" | "(m::'a::linordered_idom) \ n" | "(m::'a::linordered_idom) = n") = \K Lin_Arith.simproc\ subsection\More Inequality Reasoning\ lemma zless_add1_eq: "w < z + 1 \ w < z \ w = z" for w z :: int by arith lemma add1_zle_eq: "w + 1 \ z \ w < z" for w z :: int by arith lemma zle_diff1_eq [simp]: "w \ z - 1 \ w < z" for w z :: int by arith lemma zle_add1_eq_le [simp]: "w < z + 1 \ w \ z" for w z :: int by arith lemma int_one_le_iff_zero_less: "1 \ z \ 0 < z" for z :: int by arith lemma Ints_nonzero_abs_ge1: fixes x:: "'a :: linordered_idom" assumes "x \ Ints" "x \ 0" shows "1 \ abs x" proof (rule Ints_cases [OF \x \ Ints\]) fix z::int assume "x = of_int z" with \x \ 0\ show "1 \ \x\" apply (auto simp add: abs_if) by (metis diff_0 of_int_1 of_int_le_iff of_int_minus zle_diff1_eq) qed lemma Ints_nonzero_abs_less1: fixes x:: "'a :: linordered_idom" shows "\x \ Ints; abs x < 1\ \ x = 0" using Ints_nonzero_abs_ge1 [of x] by auto lemma Ints_eq_abs_less1: fixes x:: "'a :: linordered_idom" shows "\x \ Ints; y \ Ints\ \ x = y \ abs (x-y) < 1" using eq_iff_diff_eq_0 by (fastforce intro: Ints_nonzero_abs_less1) subsection \The functions \<^term>\nat\ and \<^term>\int\\ text \Simplify the term \<^term>\w + - z\.\ lemma one_less_nat_eq [simp]: "Suc 0 < nat z \ 1 < z" using zless_nat_conj [of 1 z] by auto lemma int_eq_iff_numeral [simp]: "int m = numeral v \ m = numeral v" by (simp add: int_eq_iff) lemma nat_abs_int_diff: "nat \int a - int b\ = (if a \ b then b - a else a - b)" by auto lemma nat_int_add: "nat (int a + int b) = a + b" by auto context ring_1 begin lemma of_int_of_nat [nitpick_simp]: "of_int k = (if k < 0 then - of_nat (nat (- k)) else of_nat (nat k))" proof (cases "k < 0") case True then have "0 \ - k" by simp then have "of_nat (nat (- k)) = of_int (- k)" by (rule of_nat_nat) with True show ?thesis by simp next case False then show ?thesis by (simp add: not_less) qed end lemma transfer_rule_of_int: + includes lifting_syntax fixes R :: "'a::ring_1 \ 'b::ring_1 \ bool" assumes [transfer_rule]: "R 0 0" "R 1 1" - "rel_fun R (rel_fun R R) plus plus" - "rel_fun R R uminus uminus" - shows "rel_fun HOL.eq R of_int of_int" + "(R ===> R ===> R) (+) (+)" + "(R ===> R) uminus uminus" + shows "((=) ===> R) of_int of_int" proof - + note assms note transfer_rule_of_nat [transfer_rule] - have [transfer_rule]: "rel_fun HOL.eq R of_nat of_nat" + have [transfer_rule]: "((=) ===> R) of_nat of_nat" by transfer_prover show ?thesis by (unfold of_int_of_nat [abs_def]) transfer_prover qed lemma nat_mult_distrib: fixes z z' :: int assumes "0 \ z" shows "nat (z * z') = nat z * nat z'" proof (cases "0 \ z'") case False with assms have "z * z' \ 0" by (simp add: not_le mult_le_0_iff) then have "nat (z * z') = 0" by simp moreover from False have "nat z' = 0" by simp ultimately show ?thesis by simp next case True with assms have ge_0: "z * z' \ 0" by (simp add: zero_le_mult_iff) show ?thesis by (rule injD [of "of_nat :: nat \ int", OF inj_of_nat]) (simp only: of_nat_mult of_nat_nat [OF True] of_nat_nat [OF assms] of_nat_nat [OF ge_0], simp) qed lemma nat_mult_distrib_neg: "z \ 0 \ nat (z * z') = nat (- z) * nat (- z')" for z z' :: int apply (rule trans) apply (rule_tac [2] nat_mult_distrib) apply auto done lemma nat_abs_mult_distrib: "nat \w * z\ = nat \w\ * nat \z\" by (cases "z = 0 \ w = 0") (auto simp add: abs_if nat_mult_distrib [symmetric] nat_mult_distrib_neg [symmetric] mult_less_0_iff) lemma int_in_range_abs [simp]: "int n \ range abs" proof (rule range_eqI) show "int n = \int n\" by simp qed lemma range_abs_Nats [simp]: "range abs = (\ :: int set)" proof - have "\k\ \ \" for k :: int by (cases k) simp_all moreover have "k \ range abs" if "k \ \" for k :: int using that by induct simp ultimately show ?thesis by blast qed lemma Suc_nat_eq_nat_zadd1: "0 \ z \ Suc (nat z) = nat (1 + z)" for z :: int by (rule sym) (simp add: nat_eq_iff) lemma diff_nat_eq_if: "nat z - nat z' = (if z' < 0 then nat z else let d = z - z' in if d < 0 then 0 else nat d)" by (simp add: Let_def nat_diff_distrib [symmetric]) lemma nat_numeral_diff_1 [simp]: "numeral v - (1::nat) = nat (numeral v - 1)" using diff_nat_numeral [of v Num.One] by simp subsection \Induction principles for int\ text \Well-founded segments of the integers.\ definition int_ge_less_than :: "int \ (int \ int) set" where "int_ge_less_than d = {(z', z). d \ z' \ z' < z}" lemma wf_int_ge_less_than: "wf (int_ge_less_than d)" proof - have "int_ge_less_than d \ measure (\z. nat (z - d))" by (auto simp add: int_ge_less_than_def) then show ?thesis by (rule wf_subset [OF wf_measure]) qed text \ This variant looks odd, but is typical of the relations suggested by RankFinder.\ definition int_ge_less_than2 :: "int \ (int \ int) set" where "int_ge_less_than2 d = {(z',z). d \ z \ z' < z}" lemma wf_int_ge_less_than2: "wf (int_ge_less_than2 d)" proof - have "int_ge_less_than2 d \ measure (\z. nat (1 + z - d))" by (auto simp add: int_ge_less_than2_def) then show ?thesis by (rule wf_subset [OF wf_measure]) qed (* `set:int': dummy construction *) theorem int_ge_induct [case_names base step, induct set: int]: fixes i :: int assumes ge: "k \ i" and base: "P k" and step: "\i. k \ i \ P i \ P (i + 1)" shows "P i" proof - have "\i::int. n = nat (i - k) \ k \ i \ P i" for n proof (induct n) case 0 then have "i = k" by arith with base show "P i" by simp next case (Suc n) then have "n = nat ((i - 1) - k)" by arith moreover have k: "k \ i - 1" using Suc.prems by arith ultimately have "P (i - 1)" by (rule Suc.hyps) from step [OF k this] show ?case by simp qed with ge show ?thesis by fast qed (* `set:int': dummy construction *) theorem int_gr_induct [case_names base step, induct set: int]: fixes i k :: int assumes gr: "k < i" and base: "P (k + 1)" and step: "\i. k < i \ P i \ P (i + 1)" shows "P i" apply (rule int_ge_induct[of "k + 1"]) using gr apply arith apply (rule base) apply (rule step) apply simp_all done theorem int_le_induct [consumes 1, case_names base step]: fixes i k :: int assumes le: "i \ k" and base: "P k" and step: "\i. i \ k \ P i \ P (i - 1)" shows "P i" proof - have "\i::int. n = nat(k-i) \ i \ k \ P i" for n proof (induct n) case 0 then have "i = k" by arith with base show "P i" by simp next case (Suc n) then have "n = nat (k - (i + 1))" by arith moreover have k: "i + 1 \ k" using Suc.prems by arith ultimately have "P (i + 1)" by (rule Suc.hyps) from step[OF k this] show ?case by simp qed with le show ?thesis by fast qed theorem int_less_induct [consumes 1, case_names base step]: fixes i k :: int assumes less: "i < k" and base: "P (k - 1)" and step: "\i. i < k \ P i \ P (i - 1)" shows "P i" apply (rule int_le_induct[of _ "k - 1"]) using less apply arith apply (rule base) apply (rule step) apply simp_all done theorem int_induct [case_names base step1 step2]: fixes k :: int assumes base: "P k" and step1: "\i. k \ i \ P i \ P (i + 1)" and step2: "\i. k \ i \ P i \ P (i - 1)" shows "P i" proof - have "i \ k \ i \ k" by arith then show ?thesis proof assume "i \ k" then show ?thesis using base by (rule int_ge_induct) (fact step1) next assume "i \ k" then show ?thesis using base by (rule int_le_induct) (fact step2) qed qed subsection \Intermediate value theorems\ lemma nat_intermed_int_val: "\i. m \ i \ i \ n \ f i = k" if "\i. m \ i \ i < n \ \f (Suc i) - f i\ \ 1" "m \ n" "f m \ k" "k \ f n" for m n :: nat and k :: int proof - have "(\if (Suc i) - f i\ \ 1) \ f 0 \ k \ k \ f n \ (\i \ n. f i = k)" for n :: nat and f apply (induct n) apply auto apply (erule_tac x = n in allE) apply (case_tac "k = f (Suc n)") apply (auto simp add: abs_if split: if_split_asm intro: le_SucI) done from this [of "n - m" "f \ plus m"] that show ?thesis apply auto apply (rule_tac x = "m + i" in exI) apply auto done qed lemma nat0_intermed_int_val: "\i\n. f i = k" if "\if (i + 1) - f i\ \ 1" "f 0 \ k" "k \ f n" for n :: nat and k :: int using nat_intermed_int_val [of 0 n f k] that by auto subsection \Products and 1, by T. M. Rasmussen\ lemma abs_zmult_eq_1: fixes m n :: int assumes mn: "\m * n\ = 1" shows "\m\ = 1" proof - from mn have 0: "m \ 0" "n \ 0" by auto have "\ 2 \ \m\" proof assume "2 \ \m\" then have "2 * \n\ \ \m\ * \n\" by (simp add: mult_mono 0) also have "\ = \m * n\" by (simp add: abs_mult) also from mn have "\ = 1" by simp finally have "2 * \n\ \ 1" . with 0 show "False" by arith qed with 0 show ?thesis by auto qed lemma pos_zmult_eq_1_iff_lemma: "m * n = 1 \ m = 1 \ m = - 1" for m n :: int using abs_zmult_eq_1 [of m n] by arith lemma pos_zmult_eq_1_iff: fixes m n :: int assumes "0 < m" shows "m * n = 1 \ m = 1 \ n = 1" proof - from assms have "m * n = 1 \ m = 1" by (auto dest: pos_zmult_eq_1_iff_lemma) then show ?thesis by (auto dest: pos_zmult_eq_1_iff_lemma) qed lemma zmult_eq_1_iff: "m * n = 1 \ (m = 1 \ n = 1) \ (m = - 1 \ n = - 1)" for m n :: int apply (rule iffI) apply (frule pos_zmult_eq_1_iff_lemma) apply (simp add: mult.commute [of m]) apply (frule pos_zmult_eq_1_iff_lemma) apply auto done lemma infinite_UNIV_int [simp]: "\ finite (UNIV::int set)" proof assume "finite (UNIV::int set)" moreover have "inj (\i::int. 2 * i)" by (rule injI) simp ultimately have "surj (\i::int. 2 * i)" by (rule finite_UNIV_inj_surj) then obtain i :: int where "1 = 2 * i" by (rule surjE) then show False by (simp add: pos_zmult_eq_1_iff) qed subsection \The divides relation\ lemma zdvd_antisym_nonneg: "0 \ m \ 0 \ n \ m dvd n \ n dvd m \ m = n" for m n :: int by (auto simp add: dvd_def mult.assoc zero_le_mult_iff zmult_eq_1_iff) lemma zdvd_antisym_abs: fixes a b :: int assumes "a dvd b" and "b dvd a" shows "\a\ = \b\" proof (cases "a = 0") case True with assms show ?thesis by simp next case False from \a dvd b\ obtain k where k: "b = a * k" unfolding dvd_def by blast from \b dvd a\ obtain k' where k': "a = b * k'" unfolding dvd_def by blast from k k' have "a = a * k * k'" by simp with mult_cancel_left1[where c="a" and b="k*k'"] have kk': "k * k' = 1" using \a \ 0\ by (simp add: mult.assoc) then have "k = 1 \ k' = 1 \ k = -1 \ k' = -1" by (simp add: zmult_eq_1_iff) with k k' show ?thesis by auto qed lemma zdvd_zdiffD: "k dvd m - n \ k dvd n \ k dvd m" for k m n :: int using dvd_add_right_iff [of k "- n" m] by simp lemma zdvd_reduce: "k dvd n + k * m \ k dvd n" for k m n :: int using dvd_add_times_triv_right_iff [of k n m] by (simp add: ac_simps) lemma dvd_imp_le_int: fixes d i :: int assumes "i \ 0" and "d dvd i" shows "\d\ \ \i\" proof - from \d dvd i\ obtain k where "i = d * k" .. with \i \ 0\ have "k \ 0" by auto then have "1 \ \k\" and "0 \ \d\" by auto then have "\d\ * 1 \ \d\ * \k\" by (rule mult_left_mono) with \i = d * k\ show ?thesis by (simp add: abs_mult) qed lemma zdvd_not_zless: fixes m n :: int assumes "0 < m" and "m < n" shows "\ n dvd m" proof from assms have "0 < n" by auto assume "n dvd m" then obtain k where k: "m = n * k" .. with \0 < m\ have "0 < n * k" by auto with \0 < n\ have "0 < k" by (simp add: zero_less_mult_iff) with k \0 < n\ \m < n\ have "n * k < n * 1" by simp with \0 < n\ \0 < k\ show False unfolding mult_less_cancel_left by auto qed lemma zdvd_mult_cancel: fixes k m n :: int assumes d: "k * m dvd k * n" and "k \ 0" shows "m dvd n" proof - from d obtain h where h: "k * n = k * m * h" unfolding dvd_def by blast have "n = m * h" proof (rule ccontr) assume "\ ?thesis" with \k \ 0\ have "k * n \ k * (m * h)" by simp with h show False by (simp add: mult.assoc) qed then show ?thesis by simp qed lemma int_dvd_int_iff [simp]: "int m dvd int n \ m dvd n" proof - have "m dvd n" if "int n = int m * k" for k proof (cases k) case (nonneg q) with that have "n = m * q" by (simp del: of_nat_mult add: of_nat_mult [symmetric]) then show ?thesis .. next case (neg q) with that have "int n = int m * (- int (Suc q))" by simp also have "\ = - (int m * int (Suc q))" by (simp only: mult_minus_right) also have "\ = - int (m * Suc q)" by (simp only: of_nat_mult [symmetric]) finally have "- int (m * Suc q) = int n" .. then show ?thesis by (simp only: negative_eq_positive) auto qed then show ?thesis by (auto simp add: dvd_def) qed lemma dvd_nat_abs_iff [simp]: "n dvd nat \k\ \ int n dvd k" proof - have "n dvd nat \k\ \ int n dvd int (nat \k\)" by (simp only: int_dvd_int_iff) then show ?thesis by simp qed lemma nat_abs_dvd_iff [simp]: "nat \k\ dvd n \ k dvd int n" proof - have "nat \k\ dvd n \ int (nat \k\) dvd int n" by (simp only: int_dvd_int_iff) then show ?thesis by simp qed lemma zdvd1_eq [simp]: "x dvd 1 \ \x\ = 1" (is "?lhs \ ?rhs") for x :: int proof assume ?lhs then have "nat \x\ dvd nat \1\" by (simp only: nat_abs_dvd_iff) simp then have "nat \x\ = 1" by simp then show ?rhs by (cases "x < 0") simp_all next assume ?rhs then have "x = 1 \ x = - 1" by auto then show ?lhs by (auto intro: dvdI) qed lemma zdvd_mult_cancel1: fixes m :: int assumes mp: "m \ 0" shows "m * n dvd m \ \n\ = 1" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (cases "n > 0") (auto simp add: minus_equation_iff) next assume ?lhs then have "m * n dvd m * 1" by simp from zdvd_mult_cancel[OF this mp] show ?rhs by (simp only: zdvd1_eq) qed lemma nat_dvd_iff: "nat z dvd m \ (if 0 \ z then z dvd int m else m = 0)" using nat_abs_dvd_iff [of z m] by (cases "z \ 0") auto lemma eq_nat_nat_iff: "0 \ z \ 0 \ z' \ nat z = nat z' \ z = z'" by (auto elim: nonneg_int_cases) lemma nat_power_eq: "0 \ z \ nat (z ^ n) = nat z ^ n" by (induct n) (simp_all add: nat_mult_distrib) lemma numeral_power_eq_nat_cancel_iff [simp]: "numeral x ^ n = nat y \ numeral x ^ n = y" using nat_eq_iff2 by auto lemma nat_eq_numeral_power_cancel_iff [simp]: "nat y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_nat_cancel_iff[of x n y] by (metis (mono_tags)) lemma numeral_power_le_nat_cancel_iff [simp]: "numeral x ^ n \ nat a \ numeral x ^ n \ a" using nat_le_eq_zle[of "numeral x ^ n" a] by (auto simp: nat_power_eq) lemma nat_le_numeral_power_cancel_iff [simp]: "nat a \ numeral x ^ n \ a \ numeral x ^ n" by (simp add: nat_le_iff) lemma numeral_power_less_nat_cancel_iff [simp]: "numeral x ^ n < nat a \ numeral x ^ n < a" using nat_less_eq_zless[of "numeral x ^ n" a] by (auto simp: nat_power_eq) lemma nat_less_numeral_power_cancel_iff [simp]: "nat a < numeral x ^ n \ a < numeral x ^ n" using nat_less_eq_zless[of a "numeral x ^ n"] by (cases "a < 0") (auto simp: nat_power_eq less_le_trans[where y=0]) lemma zdvd_imp_le: "z dvd n \ 0 < n \ z \ n" for n z :: int apply (cases n) apply auto apply (cases z) apply (auto simp add: dvd_imp_le) done lemma zdvd_period: fixes a d :: int assumes "a dvd d" shows "a dvd (x + t) \ a dvd ((x + c * d) + t)" (is "?lhs \ ?rhs") proof - from assms have "a dvd (x + t) \ a dvd ((x + t) + c * d)" by (simp add: dvd_add_left_iff) then show ?thesis by (simp add: ac_simps) qed subsection \Finiteness of intervals\ lemma finite_interval_int1 [iff]: "finite {i :: int. a \ i \ i \ b}" proof (cases "a \ b") case True then show ?thesis proof (induct b rule: int_ge_induct) case base have "{i. a \ i \ i \ a} = {a}" by auto then show ?case by simp next case (step b) then have "{i. a \ i \ i \ b + 1} = {i. a \ i \ i \ b} \ {b + 1}" by auto with step show ?case by simp qed next case False then show ?thesis by (metis (lifting, no_types) Collect_empty_eq finite.emptyI order_trans) qed lemma finite_interval_int2 [iff]: "finite {i :: int. a \ i \ i < b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto lemma finite_interval_int3 [iff]: "finite {i :: int. a < i \ i \ b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto lemma finite_interval_int4 [iff]: "finite {i :: int. a < i \ i < b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto subsection \Configuration of the code generator\ text \Constructors\ definition Pos :: "num \ int" where [simp, code_abbrev]: "Pos = numeral" definition Neg :: "num \ int" where [simp, code_abbrev]: "Neg n = - (Pos n)" code_datatype "0::int" Pos Neg text \Auxiliary operations.\ definition dup :: "int \ int" where [simp]: "dup k = k + k" lemma dup_code [code]: "dup 0 = 0" "dup (Pos n) = Pos (Num.Bit0 n)" "dup (Neg n) = Neg (Num.Bit0 n)" by (simp_all add: numeral_Bit0) definition sub :: "num \ num \ int" where [simp]: "sub m n = numeral m - numeral n" lemma sub_code [code]: "sub Num.One Num.One = 0" "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)" "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)" "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)" "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)" "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)" "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)" "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1" "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1" by (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def numeral_BitM) text \Implementations.\ lemma one_int_code [code]: "1 = Pos Num.One" by simp lemma plus_int_code [code]: "k + 0 = k" "0 + l = l" "Pos m + Pos n = Pos (m + n)" "Pos m + Neg n = sub m n" "Neg m + Pos n = sub n m" "Neg m + Neg n = Neg (m + n)" for k l :: int by simp_all lemma uminus_int_code [code]: "uminus 0 = (0::int)" "uminus (Pos m) = Neg m" "uminus (Neg m) = Pos m" by simp_all lemma minus_int_code [code]: "k - 0 = k" "0 - l = uminus l" "Pos m - Pos n = sub m n" "Pos m - Neg n = Pos (m + n)" "Neg m - Pos n = Neg (m + n)" "Neg m - Neg n = sub n m" for k l :: int by simp_all lemma times_int_code [code]: "k * 0 = 0" "0 * l = 0" "Pos m * Pos n = Pos (m * n)" "Pos m * Neg n = Neg (m * n)" "Neg m * Pos n = Neg (m * n)" "Neg m * Neg n = Pos (m * n)" for k l :: int by simp_all instantiation int :: equal begin definition "HOL.equal k l \ k = (l::int)" instance by standard (rule equal_int_def) end lemma equal_int_code [code]: "HOL.equal 0 (0::int) \ True" "HOL.equal 0 (Pos l) \ False" "HOL.equal 0 (Neg l) \ False" "HOL.equal (Pos k) 0 \ False" "HOL.equal (Pos k) (Pos l) \ HOL.equal k l" "HOL.equal (Pos k) (Neg l) \ False" "HOL.equal (Neg k) 0 \ False" "HOL.equal (Neg k) (Pos l) \ False" "HOL.equal (Neg k) (Neg l) \ HOL.equal k l" by (auto simp add: equal) lemma equal_int_refl [code nbe]: "HOL.equal k k \ True" for k :: int by (fact equal_refl) lemma less_eq_int_code [code]: "0 \ (0::int) \ True" "0 \ Pos l \ True" "0 \ Neg l \ False" "Pos k \ 0 \ False" "Pos k \ Pos l \ k \ l" "Pos k \ Neg l \ False" "Neg k \ 0 \ True" "Neg k \ Pos l \ True" "Neg k \ Neg l \ l \ k" by simp_all lemma less_int_code [code]: "0 < (0::int) \ False" "0 < Pos l \ True" "0 < Neg l \ False" "Pos k < 0 \ False" "Pos k < Pos l \ k < l" "Pos k < Neg l \ False" "Neg k < 0 \ True" "Neg k < Pos l \ True" "Neg k < Neg l \ l < k" by simp_all lemma nat_code [code]: "nat (Int.Neg k) = 0" "nat 0 = 0" "nat (Int.Pos k) = nat_of_num k" by (simp_all add: nat_of_num_numeral) lemma (in ring_1) of_int_code [code]: "of_int (Int.Neg k) = - numeral k" "of_int 0 = 0" "of_int (Int.Pos k) = numeral k" by simp_all text \Serializer setup.\ code_identifier code_module Int \ (SML) Arith and (OCaml) Arith and (Haskell) Arith quickcheck_params [default_type = int] hide_const (open) Pos Neg sub dup text \De-register \int\ as a quotient type:\ lifting_update int.lifting lifting_forget int.lifting subsection \Duplicates\ lemmas int_sum = of_nat_sum [where 'a=int] lemmas int_prod = of_nat_prod [where 'a=int] lemmas zle_int = of_nat_le_iff [where 'a=int] lemmas int_int_eq = of_nat_eq_iff [where 'a=int] lemmas nonneg_eq_int = nonneg_int_cases lemmas double_eq_0_iff = double_zero lemmas int_distrib = distrib_right [of z1 z2 w] distrib_left [of w z1 z2] left_diff_distrib [of z1 z2 w] right_diff_distrib [of w z1 z2] for z1 z2 w :: int end diff --git a/src/HOL/Lifting_Set.thy b/src/HOL/Lifting_Set.thy --- a/src/HOL/Lifting_Set.thy +++ b/src/HOL/Lifting_Set.thy @@ -1,322 +1,332 @@ (* Title: HOL/Lifting_Set.thy Author: Brian Huffman and Ondrej Kuncar *) section \Setup for Lifting/Transfer for the set type\ theory Lifting_Set imports Lifting begin subsection \Relator and predicator properties\ lemma rel_setD1: "\ rel_set R A B; x \ A \ \ \y \ B. R x y" and rel_setD2: "\ rel_set R A B; y \ B \ \ \x \ A. R x y" by (simp_all add: rel_set_def) lemma rel_set_conversep [simp]: "rel_set A\\ = (rel_set A)\\" unfolding rel_set_def by auto lemma rel_set_eq [relator_eq]: "rel_set (=) = (=)" unfolding rel_set_def fun_eq_iff by auto lemma rel_set_mono[relator_mono]: assumes "A \ B" shows "rel_set A \ rel_set B" using assms unfolding rel_set_def by blast lemma rel_set_OO[relator_distr]: "rel_set R OO rel_set S = rel_set (R OO S)" apply (rule sym) apply (intro ext) subgoal for X Z apply (rule iffI) apply (rule relcomppI [where b="{y. (\x\X. R x y) \ (\z\Z. S y z)}"]) apply (simp add: rel_set_def, fast)+ done done lemma Domainp_set[relator_domain]: "Domainp (rel_set T) = (\A. Ball A (Domainp T))" unfolding rel_set_def Domainp_iff[abs_def] apply (intro ext) apply (rule iffI) apply blast subgoal for A by (rule exI [where x="{y. \x\A. T x y}"]) fast done lemma left_total_rel_set[transfer_rule]: "left_total A \ left_total (rel_set A)" unfolding left_total_def rel_set_def apply safe subgoal for X by (rule exI [where x="{y. \x\X. A x y}"]) fast done lemma left_unique_rel_set[transfer_rule]: "left_unique A \ left_unique (rel_set A)" unfolding left_unique_def rel_set_def by fast lemma right_total_rel_set [transfer_rule]: "right_total A \ right_total (rel_set A)" using left_total_rel_set[of "A\\"] by simp lemma right_unique_rel_set [transfer_rule]: "right_unique A \ right_unique (rel_set A)" unfolding right_unique_def rel_set_def by fast lemma bi_total_rel_set [transfer_rule]: "bi_total A \ bi_total (rel_set A)" by(simp add: bi_total_alt_def left_total_rel_set right_total_rel_set) lemma bi_unique_rel_set [transfer_rule]: "bi_unique A \ bi_unique (rel_set A)" unfolding bi_unique_def rel_set_def by fast lemma set_relator_eq_onp [relator_eq_onp]: "rel_set (eq_onp P) = eq_onp (\A. Ball A P)" unfolding fun_eq_iff rel_set_def eq_onp_def Ball_def by fast lemma bi_unique_rel_set_lemma: assumes "bi_unique R" and "rel_set R X Y" obtains f where "Y = image f X" and "inj_on f X" and "\x\X. R x (f x)" proof define f where "f x = (THE y. R x y)" for x { fix x assume "x \ X" with \rel_set R X Y\ \bi_unique R\ have "R x (f x)" by (simp add: bi_unique_def rel_set_def f_def) (metis theI) with assms \x \ X\ have "R x (f x)" "\x'\X. R x' (f x) \ x = x'" "\y\Y. R x y \ y = f x" "f x \ Y" by (fastforce simp add: bi_unique_def rel_set_def)+ } note * = this moreover { fix y assume "y \ Y" with \rel_set R X Y\ *(3) \y \ Y\ have "\x\X. y = f x" by (fastforce simp: rel_set_def) } ultimately show "\x\X. R x (f x)" "Y = image f X" "inj_on f X" by (auto simp: inj_on_def image_iff) qed subsection \Quotient theorem for the Lifting package\ lemma Quotient_set[quot_map]: assumes "Quotient R Abs Rep T" shows "Quotient (rel_set R) (image Abs) (image Rep) (rel_set T)" using assms unfolding Quotient_alt_def4 apply (simp add: rel_set_OO[symmetric]) apply (simp add: rel_set_def) apply fast done subsection \Transfer rules for the Transfer package\ subsubsection \Unconditional transfer rules\ context includes lifting_syntax begin lemma empty_transfer [transfer_rule]: "(rel_set A) {} {}" unfolding rel_set_def by simp lemma insert_transfer [transfer_rule]: "(A ===> rel_set A ===> rel_set A) insert insert" unfolding rel_fun_def rel_set_def by auto lemma union_transfer [transfer_rule]: "(rel_set A ===> rel_set A ===> rel_set A) union union" unfolding rel_fun_def rel_set_def by auto lemma Union_transfer [transfer_rule]: "(rel_set (rel_set A) ===> rel_set A) Union Union" unfolding rel_fun_def rel_set_def by simp fast lemma image_transfer [transfer_rule]: "((A ===> B) ===> rel_set A ===> rel_set B) image image" unfolding rel_fun_def rel_set_def by simp fast lemma UNION_transfer [transfer_rule]: \ \TODO deletion candidate\ "(rel_set A ===> (A ===> rel_set B) ===> rel_set B) (\A f. \(f ` A)) (\A f. \(f ` A))" by transfer_prover lemma Ball_transfer [transfer_rule]: "(rel_set A ===> (A ===> (=)) ===> (=)) Ball Ball" unfolding rel_set_def rel_fun_def by fast lemma Bex_transfer [transfer_rule]: "(rel_set A ===> (A ===> (=)) ===> (=)) Bex Bex" unfolding rel_set_def rel_fun_def by fast lemma Pow_transfer [transfer_rule]: "(rel_set A ===> rel_set (rel_set A)) Pow Pow" apply (rule rel_funI) apply (rule rel_setI) subgoal for X Y X' apply (rule rev_bexI [where x="{y\Y. \x\X'. A x y}"]) apply clarsimp apply (simp add: rel_set_def) apply fast done subgoal for X Y Y' apply (rule rev_bexI [where x="{x\X. \y\Y'. A x y}"]) apply clarsimp apply (simp add: rel_set_def) apply fast done done lemma rel_set_transfer [transfer_rule]: "((A ===> B ===> (=)) ===> rel_set A ===> rel_set B ===> (=)) rel_set rel_set" unfolding rel_fun_def rel_set_def by fast lemma bind_transfer [transfer_rule]: "(rel_set A ===> (A ===> rel_set B) ===> rel_set B) Set.bind Set.bind" unfolding bind_UNION [abs_def] by transfer_prover lemma INF_parametric [transfer_rule]: \ \TODO deletion candidate\ "(rel_set A ===> (A ===> HOL.eq) ===> HOL.eq) (\A f. Inf (f ` A)) (\A f. Inf (f ` A))" by transfer_prover lemma SUP_parametric [transfer_rule]: \ \TODO deletion candidate\ "(rel_set R ===> (R ===> HOL.eq) ===> HOL.eq) (\A f. Sup (f ` A)) (\A f. Sup (f ` A))" by transfer_prover subsubsection \Rules requiring bi-unique, bi-total or right-total relations\ lemma member_transfer [transfer_rule]: assumes "bi_unique A" shows "(A ===> rel_set A ===> (=)) (\) (\)" using assms unfolding rel_fun_def rel_set_def bi_unique_def by fast lemma right_total_Collect_transfer[transfer_rule]: assumes "right_total A" shows "((A ===> (=)) ===> rel_set A) (\P. Collect (\x. P x \ Domainp A x)) Collect" using assms unfolding right_total_def rel_set_def rel_fun_def Domainp_iff by fast lemma Collect_transfer [transfer_rule]: assumes "bi_total A" shows "((A ===> (=)) ===> rel_set A) Collect Collect" using assms unfolding rel_fun_def rel_set_def bi_total_def by fast lemma inter_transfer [transfer_rule]: assumes "bi_unique A" shows "(rel_set A ===> rel_set A ===> rel_set A) inter inter" using assms unfolding rel_fun_def rel_set_def bi_unique_def by fast lemma Diff_transfer [transfer_rule]: assumes "bi_unique A" shows "(rel_set A ===> rel_set A ===> rel_set A) (-) (-)" using assms unfolding rel_fun_def rel_set_def bi_unique_def unfolding Ball_def Bex_def Diff_eq by (safe, simp, metis, simp, metis) lemma subset_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(rel_set A ===> rel_set A ===> (=)) (\) (\)" unfolding subset_eq [abs_def] by transfer_prover +context + includes lifting_syntax +begin + lemma strict_subset_transfer [transfer_rule]: - includes lifting_syntax assumes [transfer_rule]: "bi_unique A" shows "(rel_set A ===> rel_set A ===> (=)) (\) (\)" unfolding subset_not_subset_eq by transfer_prover +end + declare right_total_UNIV_transfer[transfer_rule] lemma UNIV_transfer [transfer_rule]: assumes "bi_total A" shows "(rel_set A) UNIV UNIV" using assms unfolding rel_set_def bi_total_def by simp lemma right_total_Compl_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" and [transfer_rule]: "right_total A" shows "(rel_set A ===> rel_set A) (\S. uminus S \ Collect (Domainp A)) uminus" unfolding Compl_eq [abs_def] by (subst Collect_conj_eq[symmetric]) transfer_prover lemma Compl_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" and [transfer_rule]: "bi_total A" shows "(rel_set A ===> rel_set A) uminus uminus" unfolding Compl_eq [abs_def] by transfer_prover lemma right_total_Inter_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" and [transfer_rule]: "right_total A" shows "(rel_set (rel_set A) ===> rel_set A) (\S. \S \ Collect (Domainp A)) Inter" unfolding Inter_eq[abs_def] by (subst Collect_conj_eq[symmetric]) transfer_prover lemma Inter_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" and [transfer_rule]: "bi_total A" shows "(rel_set (rel_set A) ===> rel_set A) Inter Inter" unfolding Inter_eq [abs_def] by transfer_prover lemma filter_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "((A ===> (=)) ===> rel_set A ===> rel_set A) Set.filter Set.filter" unfolding Set.filter_def[abs_def] rel_fun_def rel_set_def by blast lemma finite_transfer [transfer_rule]: "bi_unique A \ (rel_set A ===> (=)) finite finite" by (rule rel_funI, erule (1) bi_unique_rel_set_lemma) (auto dest: finite_imageD) lemma card_transfer [transfer_rule]: "bi_unique A \ (rel_set A ===> (=)) card card" by (rule rel_funI, erule (1) bi_unique_rel_set_lemma) (simp add: card_image) +context + includes lifting_syntax +begin + lemma vimage_right_total_transfer[transfer_rule]: - includes lifting_syntax assumes [transfer_rule]: "bi_unique B" "right_total A" shows "((A ===> B) ===> rel_set B ===> rel_set A) (\f X. f -` X \ Collect (Domainp A)) vimage" proof - let ?vimage = "(\f B. {x. f x \ B \ Domainp A x})" have "((A ===> B) ===> rel_set B ===> rel_set A) ?vimage vimage" unfolding vimage_def by transfer_prover also have "?vimage = (\f X. f -` X \ Collect (Domainp A))" by auto finally show ?thesis . qed +end + lemma vimage_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" "bi_unique B" shows "((A ===> B) ===> rel_set B ===> rel_set A) vimage vimage" unfolding vimage_def[abs_def] by transfer_prover lemma Image_parametric [transfer_rule]: assumes "bi_unique A" shows "(rel_set (rel_prod A B) ===> rel_set A ===> rel_set B) (``) (``)" by (intro rel_funI rel_setI) (force dest: rel_setD1 bi_uniqueDr[OF assms], force dest: rel_setD2 bi_uniqueDl[OF assms]) lemma inj_on_transfer[transfer_rule]: "((A ===> B) ===> rel_set A ===> (=)) inj_on inj_on" if [transfer_rule]: "bi_unique A" "bi_unique B" unfolding inj_on_def by transfer_prover end lemma (in comm_monoid_set) F_parametric [transfer_rule]: fixes A :: "'b \ 'c \ bool" assumes "bi_unique A" shows "rel_fun (rel_fun A (=)) (rel_fun (rel_set A) (=)) F F" proof (rule rel_funI)+ fix f :: "'b \ 'a" and g S T assume "rel_fun A (=) f g" "rel_set A S T" with \bi_unique A\ obtain i where "bij_betw i S T" "\x. x \ S \ f x = g (i x)" by (auto elim: bi_unique_rel_set_lemma simp: rel_fun_def bij_betw_def) then show "F f S = F g T" by (simp add: reindex_bij_betw) qed lemmas sum_parametric = sum.F_parametric lemmas prod_parametric = prod.F_parametric lemma rel_set_UNION: assumes [transfer_rule]: "rel_set Q A B" "rel_fun Q (rel_set R) f g" shows "rel_set R (\(f ` A)) (\(g ` B))" by transfer_prover end diff --git a/src/HOL/Num.thy b/src/HOL/Num.thy --- a/src/HOL/Num.thy +++ b/src/HOL/Num.thy @@ -1,1427 +1,1435 @@ (* Title: HOL/Num.thy Author: Florian Haftmann Author: Brian Huffman *) section \Binary Numerals\ theory Num imports BNF_Least_Fixpoint Transfer begin subsection \The \num\ type\ datatype num = One | Bit0 num | Bit1 num text \Increment function for type \<^typ>\num\\ primrec inc :: "num \ num" where "inc One = Bit0 One" | "inc (Bit0 x) = Bit1 x" | "inc (Bit1 x) = Bit0 (inc x)" text \Converting between type \<^typ>\num\ and type \<^typ>\nat\\ primrec nat_of_num :: "num \ nat" where "nat_of_num One = Suc 0" | "nat_of_num (Bit0 x) = nat_of_num x + nat_of_num x" | "nat_of_num (Bit1 x) = Suc (nat_of_num x + nat_of_num x)" primrec num_of_nat :: "nat \ num" where "num_of_nat 0 = One" | "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)" lemma nat_of_num_pos: "0 < nat_of_num x" by (induct x) simp_all lemma nat_of_num_neq_0: " nat_of_num x \ 0" by (induct x) simp_all lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)" by (induct x) simp_all lemma num_of_nat_double: "0 < n \ num_of_nat (n + n) = Bit0 (num_of_nat n)" by (induct n) simp_all text \Type \<^typ>\num\ is isomorphic to the strictly positive natural numbers.\ lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x" by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos) lemma num_of_nat_inverse: "0 < n \ nat_of_num (num_of_nat n) = n" by (induct n) (simp_all add: nat_of_num_inc) lemma num_eq_iff: "x = y \ nat_of_num x = nat_of_num y" apply safe apply (drule arg_cong [where f=num_of_nat]) apply (simp add: nat_of_num_inverse) done lemma num_induct [case_names One inc]: fixes P :: "num \ bool" assumes One: "P One" and inc: "\x. P x \ P (inc x)" shows "P x" proof - obtain n where n: "Suc n = nat_of_num x" by (cases "nat_of_num x") (simp_all add: nat_of_num_neq_0) have "P (num_of_nat (Suc n))" proof (induct n) case 0 from One show ?case by simp next case (Suc n) then have "P (inc (num_of_nat (Suc n)))" by (rule inc) then show "P (num_of_nat (Suc (Suc n)))" by simp qed with n show "P x" by (simp add: nat_of_num_inverse) qed text \ From now on, there are two possible models for \<^typ>\num\: as positive naturals (rule \num_induct\) and as digit representation (rules \num.induct\, \num.cases\). \ subsection \Numeral operations\ instantiation num :: "{plus,times,linorder}" begin definition [code del]: "m + n = num_of_nat (nat_of_num m + nat_of_num n)" definition [code del]: "m * n = num_of_nat (nat_of_num m * nat_of_num n)" definition [code del]: "m \ n \ nat_of_num m \ nat_of_num n" definition [code del]: "m < n \ nat_of_num m < nat_of_num n" instance by standard (auto simp add: less_num_def less_eq_num_def num_eq_iff) end lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y" unfolding plus_num_def by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos) lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y" unfolding times_num_def by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos) lemma add_num_simps [simp, code]: "One + One = Bit0 One" "One + Bit0 n = Bit1 n" "One + Bit1 n = Bit0 (n + One)" "Bit0 m + One = Bit1 m" "Bit0 m + Bit0 n = Bit0 (m + n)" "Bit0 m + Bit1 n = Bit1 (m + n)" "Bit1 m + One = Bit0 (m + One)" "Bit1 m + Bit0 n = Bit1 (m + n)" "Bit1 m + Bit1 n = Bit0 (m + n + One)" by (simp_all add: num_eq_iff nat_of_num_add) lemma mult_num_simps [simp, code]: "m * One = m" "One * n = n" "Bit0 m * Bit0 n = Bit0 (Bit0 (m * n))" "Bit0 m * Bit1 n = Bit0 (m * Bit1 n)" "Bit1 m * Bit0 n = Bit0 (Bit1 m * n)" "Bit1 m * Bit1 n = Bit1 (m + n + Bit0 (m * n))" by (simp_all add: num_eq_iff nat_of_num_add nat_of_num_mult distrib_right distrib_left) lemma eq_num_simps: "One = One \ True" "One = Bit0 n \ False" "One = Bit1 n \ False" "Bit0 m = One \ False" "Bit1 m = One \ False" "Bit0 m = Bit0 n \ m = n" "Bit0 m = Bit1 n \ False" "Bit1 m = Bit0 n \ False" "Bit1 m = Bit1 n \ m = n" by simp_all lemma le_num_simps [simp, code]: "One \ n \ True" "Bit0 m \ One \ False" "Bit1 m \ One \ False" "Bit0 m \ Bit0 n \ m \ n" "Bit0 m \ Bit1 n \ m \ n" "Bit1 m \ Bit1 n \ m \ n" "Bit1 m \ Bit0 n \ m < n" using nat_of_num_pos [of n] nat_of_num_pos [of m] by (auto simp add: less_eq_num_def less_num_def) lemma less_num_simps [simp, code]: "m < One \ False" "One < Bit0 n \ True" "One < Bit1 n \ True" "Bit0 m < Bit0 n \ m < n" "Bit0 m < Bit1 n \ m \ n" "Bit1 m < Bit1 n \ m < n" "Bit1 m < Bit0 n \ m < n" using nat_of_num_pos [of n] nat_of_num_pos [of m] by (auto simp add: less_eq_num_def less_num_def) lemma le_num_One_iff: "x \ num.One \ x = num.One" by (simp add: antisym_conv) text \Rules using \One\ and \inc\ as constructors.\ lemma add_One: "x + One = inc x" by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) lemma add_One_commute: "One + n = n + One" by (induct n) simp_all lemma add_inc: "x + inc y = inc (x + y)" by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) lemma mult_inc: "x * inc y = x * y + x" by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc) text \The \<^const>\num_of_nat\ conversion.\ lemma num_of_nat_One: "n \ 1 \ num_of_nat n = One" by (cases n) simp_all lemma num_of_nat_plus_distrib: "0 < m \ 0 < n \ num_of_nat (m + n) = num_of_nat m + num_of_nat n" by (induct n) (auto simp add: add_One add_One_commute add_inc) text \A double-and-decrement function.\ primrec BitM :: "num \ num" where "BitM One = One" | "BitM (Bit0 n) = Bit1 (BitM n)" | "BitM (Bit1 n) = Bit1 (Bit0 n)" lemma BitM_plus_one: "BitM n + One = Bit0 n" by (induct n) simp_all lemma one_plus_BitM: "One + BitM n = Bit0 n" unfolding add_One_commute BitM_plus_one .. text \Squaring and exponentiation.\ primrec sqr :: "num \ num" where "sqr One = One" | "sqr (Bit0 n) = Bit0 (Bit0 (sqr n))" | "sqr (Bit1 n) = Bit1 (Bit0 (sqr n + n))" primrec pow :: "num \ num \ num" where "pow x One = x" | "pow x (Bit0 y) = sqr (pow x y)" | "pow x (Bit1 y) = sqr (pow x y) * x" lemma nat_of_num_sqr: "nat_of_num (sqr x) = nat_of_num x * nat_of_num x" by (induct x) (simp_all add: algebra_simps nat_of_num_add) lemma sqr_conv_mult: "sqr x = x * x" by (simp add: num_eq_iff nat_of_num_sqr nat_of_num_mult) lemma num_double [simp]: "num.Bit0 num.One * n = num.Bit0 n" by (simp add: num_eq_iff nat_of_num_mult) subsection \Binary numerals\ text \ We embed binary representations into a generic algebraic structure using \numeral\. \ class numeral = one + semigroup_add begin primrec numeral :: "num \ 'a" where numeral_One: "numeral One = 1" | numeral_Bit0: "numeral (Bit0 n) = numeral n + numeral n" | numeral_Bit1: "numeral (Bit1 n) = numeral n + numeral n + 1" lemma numeral_code [code]: "numeral One = 1" "numeral (Bit0 n) = (let m = numeral n in m + m)" "numeral (Bit1 n) = (let m = numeral n in m + m + 1)" by (simp_all add: Let_def) lemma one_plus_numeral_commute: "1 + numeral x = numeral x + 1" proof (induct x) case One then show ?case by simp next case Bit0 then show ?case by (simp add: add.assoc [symmetric]) (simp add: add.assoc) next case Bit1 then show ?case by (simp add: add.assoc [symmetric]) (simp add: add.assoc) qed lemma numeral_inc: "numeral (inc x) = numeral x + 1" proof (induct x) case One then show ?case by simp next case Bit0 then show ?case by simp next case (Bit1 x) have "numeral x + (1 + numeral x) + 1 = numeral x + (numeral x + 1) + 1" by (simp only: one_plus_numeral_commute) with Bit1 show ?case by (simp add: add.assoc) qed declare numeral.simps [simp del] abbreviation "Numeral1 \ numeral One" declare numeral_One [code_post] end text \Numeral syntax.\ syntax "_Numeral" :: "num_const \ 'a" ("_") ML_file \Tools/numeral.ML\ parse_translation \ let fun numeral_tr [(c as Const (\<^syntax_const>\_constrain\, _)) $ t $ u] = c $ numeral_tr [t] $ u | numeral_tr [Const (num, _)] = (Numeral.mk_number_syntax o #value o Lexicon.read_num) num | numeral_tr ts = raise TERM ("numeral_tr", ts); in [(\<^syntax_const>\_Numeral\, K numeral_tr)] end \ typed_print_translation \ let fun num_tr' ctxt T [n] = let val k = Numeral.dest_num_syntax n; val t' = Syntax.const \<^syntax_const>\_Numeral\ $ Syntax.free (string_of_int k); in (case T of Type (\<^type_name>\fun\, [_, T']) => if Printer.type_emphasis ctxt T' then Syntax.const \<^syntax_const>\_constrain\ $ t' $ Syntax_Phases.term_of_typ ctxt T' else t' | _ => if T = dummyT then t' else raise Match) end; in [(\<^const_syntax>\numeral\, num_tr')] end \ subsection \Class-specific numeral rules\ text \\<^const>\numeral\ is a morphism.\ subsubsection \Structures with addition: class \numeral\\ context numeral begin lemma numeral_add: "numeral (m + n) = numeral m + numeral n" by (induct n rule: num_induct) (simp_all only: numeral_One add_One add_inc numeral_inc add.assoc) lemma numeral_plus_numeral: "numeral m + numeral n = numeral (m + n)" by (rule numeral_add [symmetric]) lemma numeral_plus_one: "numeral n + 1 = numeral (n + One)" using numeral_add [of n One] by (simp add: numeral_One) lemma one_plus_numeral: "1 + numeral n = numeral (One + n)" using numeral_add [of One n] by (simp add: numeral_One) lemma one_add_one: "1 + 1 = 2" using numeral_add [of One One] by (simp add: numeral_One) lemmas add_numeral_special = numeral_plus_one one_plus_numeral one_add_one end subsubsection \Structures with negation: class \neg_numeral\\ class neg_numeral = numeral + group_add begin lemma uminus_numeral_One: "- Numeral1 = - 1" by (simp add: numeral_One) text \Numerals form an abelian subgroup.\ inductive is_num :: "'a \ bool" where "is_num 1" | "is_num x \ is_num (- x)" | "is_num x \ is_num y \ is_num (x + y)" lemma is_num_numeral: "is_num (numeral k)" by (induct k) (simp_all add: numeral.simps is_num.intros) lemma is_num_add_commute: "is_num x \ is_num y \ x + y = y + x" apply (induct x rule: is_num.induct) apply (induct y rule: is_num.induct) apply simp apply (rule_tac a=x in add_left_imp_eq) apply (rule_tac a=x in add_right_imp_eq) apply (simp add: add.assoc) apply (simp add: add.assoc [symmetric]) apply (simp add: add.assoc) apply (rule_tac a=x in add_left_imp_eq) apply (rule_tac a=x in add_right_imp_eq) apply (simp add: add.assoc) apply (simp add: add.assoc) apply (simp add: add.assoc [symmetric]) done lemma is_num_add_left_commute: "is_num x \ is_num y \ x + (y + z) = y + (x + z)" by (simp only: add.assoc [symmetric] is_num_add_commute) lemmas is_num_normalize = add.assoc is_num_add_commute is_num_add_left_commute is_num.intros is_num_numeral minus_add definition dbl :: "'a \ 'a" where "dbl x = x + x" definition dbl_inc :: "'a \ 'a" where "dbl_inc x = x + x + 1" definition dbl_dec :: "'a \ 'a" where "dbl_dec x = x + x - 1" definition sub :: "num \ num \ 'a" where "sub k l = numeral k - numeral l" lemma numeral_BitM: "numeral (BitM n) = numeral (Bit0 n) - 1" by (simp only: BitM_plus_one [symmetric] numeral_add numeral_One eq_diff_eq) lemma dbl_simps [simp]: "dbl (- numeral k) = - dbl (numeral k)" "dbl 0 = 0" "dbl 1 = 2" "dbl (- 1) = - 2" "dbl (numeral k) = numeral (Bit0 k)" by (simp_all add: dbl_def numeral.simps minus_add) lemma dbl_inc_simps [simp]: "dbl_inc (- numeral k) = - dbl_dec (numeral k)" "dbl_inc 0 = 1" "dbl_inc 1 = 3" "dbl_inc (- 1) = - 1" "dbl_inc (numeral k) = numeral (Bit1 k)" by (simp_all add: dbl_inc_def dbl_dec_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff) lemma dbl_dec_simps [simp]: "dbl_dec (- numeral k) = - dbl_inc (numeral k)" "dbl_dec 0 = - 1" "dbl_dec 1 = 1" "dbl_dec (- 1) = - 3" "dbl_dec (numeral k) = numeral (BitM k)" by (simp_all add: dbl_dec_def dbl_inc_def numeral.simps numeral_BitM is_num_normalize) lemma sub_num_simps [simp]: "sub One One = 0" "sub One (Bit0 l) = - numeral (BitM l)" "sub One (Bit1 l) = - numeral (Bit0 l)" "sub (Bit0 k) One = numeral (BitM k)" "sub (Bit1 k) One = numeral (Bit0 k)" "sub (Bit0 k) (Bit0 l) = dbl (sub k l)" "sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)" "sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)" "sub (Bit1 k) (Bit1 l) = dbl (sub k l)" by (simp_all add: dbl_def dbl_dec_def dbl_inc_def sub_def numeral.simps numeral_BitM is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma add_neg_numeral_simps: "numeral m + - numeral n = sub m n" "- numeral m + numeral n = sub n m" "- numeral m + - numeral n = - (numeral m + numeral n)" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma add_neg_numeral_special: "1 + - numeral m = sub One m" "- numeral m + 1 = sub One m" "numeral m + - 1 = sub m One" "- 1 + numeral n = sub n One" "- 1 + - numeral n = - numeral (inc n)" "- numeral m + - 1 = - numeral (inc m)" "1 + - 1 = 0" "- 1 + 1 = 0" "- 1 + - 1 = - 2" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize right_minus numeral_inc del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma diff_numeral_simps: "numeral m - numeral n = sub m n" "numeral m - - numeral n = numeral (m + n)" "- numeral m - numeral n = - numeral (m + n)" "- numeral m - - numeral n = sub n m" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma diff_numeral_special: "1 - numeral n = sub One n" "numeral m - 1 = sub m One" "1 - - numeral n = numeral (One + n)" "- numeral m - 1 = - numeral (m + One)" "- 1 - numeral n = - numeral (inc n)" "numeral m - - 1 = numeral (inc m)" "- 1 - - numeral n = sub n One" "- numeral m - - 1 = sub One m" "1 - 1 = 0" "- 1 - 1 = - 2" "1 - - 1 = 2" "- 1 - - 1 = 0" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize numeral_inc del: add_uminus_conv_diff add: diff_conv_add_uminus) end subsubsection \Structures with multiplication: class \semiring_numeral\\ class semiring_numeral = semiring + monoid_mult begin subclass numeral .. lemma numeral_mult: "numeral (m * n) = numeral m * numeral n" by (induct n rule: num_induct) (simp_all add: numeral_One mult_inc numeral_inc numeral_add distrib_left) lemma numeral_times_numeral: "numeral m * numeral n = numeral (m * n)" by (rule numeral_mult [symmetric]) lemma mult_2: "2 * z = z + z" by (simp add: one_add_one [symmetric] distrib_right) lemma mult_2_right: "z * 2 = z + z" by (simp add: one_add_one [symmetric] distrib_left) lemma left_add_twice: "a + (a + b) = 2 * a + b" by (simp add: mult_2 ac_simps) end subsubsection \Structures with a zero: class \semiring_1\\ context semiring_1 begin subclass semiring_numeral .. lemma of_nat_numeral [simp]: "of_nat (numeral n) = numeral n" by (induct n) (simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1) lemma numeral_unfold_funpow: "numeral k = ((+) 1 ^^ numeral k) 0" unfolding of_nat_def [symmetric] by simp end +context + includes lifting_syntax +begin + lemma transfer_rule_numeral: - fixes R :: "'a::semiring_1 \ 'b::semiring_1 \ bool" - assumes [transfer_rule]: "R 0 0" "R 1 1" - "rel_fun R (rel_fun R R) plus plus" - shows "rel_fun HOL.eq R numeral numeral" - apply (subst (2) numeral_unfold_funpow [abs_def]) - apply (subst (1) numeral_unfold_funpow [abs_def]) - apply transfer_prover - done + "((=) ===> R) numeral numeral" + if [transfer_rule]: "R 0 0" "R 1 1" + "(R ===> R ===> R) (+) (+)" + for R :: "'a::semiring_1 \ 'b::semiring_1 \ bool" +proof - + have "((=) ===> R) (\k. ((+) 1 ^^ numeral k) 0) (\k. ((+) 1 ^^ numeral k) 0)" + by transfer_prover + then show ?thesis + by (simp flip: numeral_unfold_funpow [abs_def]) +qed + +end lemma nat_of_num_numeral [code_abbrev]: "nat_of_num = numeral" proof fix n have "numeral n = nat_of_num n" by (induct n) (simp_all add: numeral.simps) then show "nat_of_num n = numeral n" by simp qed lemma nat_of_num_code [code]: "nat_of_num One = 1" "nat_of_num (Bit0 n) = (let m = nat_of_num n in m + m)" "nat_of_num (Bit1 n) = (let m = nat_of_num n in Suc (m + m))" by (simp_all add: Let_def) subsubsection \Equality: class \semiring_char_0\\ context semiring_char_0 begin lemma numeral_eq_iff: "numeral m = numeral n \ m = n" by (simp only: of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] of_nat_eq_iff num_eq_iff) lemma numeral_eq_one_iff: "numeral n = 1 \ n = One" by (rule numeral_eq_iff [of n One, unfolded numeral_One]) lemma one_eq_numeral_iff: "1 = numeral n \ One = n" by (rule numeral_eq_iff [of One n, unfolded numeral_One]) lemma numeral_neq_zero: "numeral n \ 0" by (simp add: of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] nat_of_num_pos) lemma zero_neq_numeral: "0 \ numeral n" unfolding eq_commute [of 0] by (rule numeral_neq_zero) lemmas eq_numeral_simps [simp] = numeral_eq_iff numeral_eq_one_iff one_eq_numeral_iff numeral_neq_zero zero_neq_numeral end subsubsection \Comparisons: class \linordered_nonzero_semiring\\ context linordered_nonzero_semiring begin lemma numeral_le_iff: "numeral m \ numeral n \ m \ n" proof - have "of_nat (numeral m) \ of_nat (numeral n) \ m \ n" by (simp only: less_eq_num_def nat_of_num_numeral of_nat_le_iff) then show ?thesis by simp qed lemma one_le_numeral: "1 \ numeral n" using numeral_le_iff [of num.One n] by (simp add: numeral_One) lemma numeral_le_one_iff: "numeral n \ 1 \ n \ num.One" using numeral_le_iff [of n num.One] by (simp add: numeral_One) lemma numeral_less_iff: "numeral m < numeral n \ m < n" proof - have "of_nat (numeral m) < of_nat (numeral n) \ m < n" unfolding less_num_def nat_of_num_numeral of_nat_less_iff .. then show ?thesis by simp qed lemma not_numeral_less_one: "\ numeral n < 1" using numeral_less_iff [of n num.One] by (simp add: numeral_One) lemma one_less_numeral_iff: "1 < numeral n \ num.One < n" using numeral_less_iff [of num.One n] by (simp add: numeral_One) lemma zero_le_numeral: "0 \ numeral n" using dual_order.trans one_le_numeral zero_le_one by blast lemma zero_less_numeral: "0 < numeral n" using less_linear not_numeral_less_one order.strict_trans zero_less_one by blast lemma not_numeral_le_zero: "\ numeral n \ 0" by (simp add: not_le zero_less_numeral) lemma not_numeral_less_zero: "\ numeral n < 0" by (simp add: not_less zero_le_numeral) lemmas le_numeral_extra = zero_le_one not_one_le_zero order_refl [of 0] order_refl [of 1] lemmas less_numeral_extra = zero_less_one not_one_less_zero less_irrefl [of 0] less_irrefl [of 1] lemmas le_numeral_simps [simp] = numeral_le_iff one_le_numeral numeral_le_one_iff zero_le_numeral not_numeral_le_zero lemmas less_numeral_simps [simp] = numeral_less_iff one_less_numeral_iff not_numeral_less_one zero_less_numeral not_numeral_less_zero lemma min_0_1 [simp]: fixes min' :: "'a \ 'a \ 'a" defines "min' \ min" shows "min' 0 1 = 0" "min' 1 0 = 0" "min' 0 (numeral x) = 0" "min' (numeral x) 0 = 0" "min' 1 (numeral x) = 1" "min' (numeral x) 1 = 1" by (simp_all add: min'_def min_def le_num_One_iff) lemma max_0_1 [simp]: fixes max' :: "'a \ 'a \ 'a" defines "max' \ max" shows "max' 0 1 = 1" "max' 1 0 = 1" "max' 0 (numeral x) = numeral x" "max' (numeral x) 0 = numeral x" "max' 1 (numeral x) = numeral x" "max' (numeral x) 1 = numeral x" by (simp_all add: max'_def max_def le_num_One_iff) end text \Unfold \min\ and \max\ on numerals.\ lemmas max_number_of [simp] = max_def [of "numeral u" "numeral v"] max_def [of "numeral u" "- numeral v"] max_def [of "- numeral u" "numeral v"] max_def [of "- numeral u" "- numeral v"] for u v lemmas min_number_of [simp] = min_def [of "numeral u" "numeral v"] min_def [of "numeral u" "- numeral v"] min_def [of "- numeral u" "numeral v"] min_def [of "- numeral u" "- numeral v"] for u v subsubsection \Multiplication and negation: class \ring_1\\ context ring_1 begin subclass neg_numeral .. lemma mult_neg_numeral_simps: "- numeral m * - numeral n = numeral (m * n)" "- numeral m * numeral n = - numeral (m * n)" "numeral m * - numeral n = - numeral (m * n)" by (simp_all only: mult_minus_left mult_minus_right minus_minus numeral_mult) lemma mult_minus1 [simp]: "- 1 * z = - z" by (simp add: numeral.simps) lemma mult_minus1_right [simp]: "z * - 1 = - z" by (simp add: numeral.simps) end subsubsection \Equality using \iszero\ for rings with non-zero characteristic\ context ring_1 begin definition iszero :: "'a \ bool" where "iszero z \ z = 0" lemma iszero_0 [simp]: "iszero 0" by (simp add: iszero_def) lemma not_iszero_1 [simp]: "\ iszero 1" by (simp add: iszero_def) lemma not_iszero_Numeral1: "\ iszero Numeral1" by (simp add: numeral_One) lemma not_iszero_neg_1 [simp]: "\ iszero (- 1)" by (simp add: iszero_def) lemma not_iszero_neg_Numeral1: "\ iszero (- Numeral1)" by (simp add: numeral_One) lemma iszero_neg_numeral [simp]: "iszero (- numeral w) \ iszero (numeral w)" unfolding iszero_def by (rule neg_equal_0_iff_equal) lemma eq_iff_iszero_diff: "x = y \ iszero (x - y)" unfolding iszero_def by (rule eq_iff_diff_eq_0) text \ The \eq_numeral_iff_iszero\ lemmas are not declared \[simp]\ by default, because for rings of characteristic zero, better simp rules are possible. For a type like integers mod \n\, type-instantiated versions of these rules should be added to the simplifier, along with a type-specific rule for deciding propositions of the form \iszero (numeral w)\. bh: Maybe it would not be so bad to just declare these as simp rules anyway? I should test whether these rules take precedence over the \ring_char_0\ rules in the simplifier. \ lemma eq_numeral_iff_iszero: "numeral x = numeral y \ iszero (sub x y)" "numeral x = - numeral y \ iszero (numeral (x + y))" "- numeral x = numeral y \ iszero (numeral (x + y))" "- numeral x = - numeral y \ iszero (sub y x)" "numeral x = 1 \ iszero (sub x One)" "1 = numeral y \ iszero (sub One y)" "- numeral x = 1 \ iszero (numeral (x + One))" "1 = - numeral y \ iszero (numeral (One + y))" "numeral x = 0 \ iszero (numeral x)" "0 = numeral y \ iszero (numeral y)" "- numeral x = 0 \ iszero (numeral x)" "0 = - numeral y \ iszero (numeral y)" unfolding eq_iff_iszero_diff diff_numeral_simps diff_numeral_special by simp_all end subsubsection \Equality and negation: class \ring_char_0\\ context ring_char_0 begin lemma not_iszero_numeral [simp]: "\ iszero (numeral w)" by (simp add: iszero_def) lemma neg_numeral_eq_iff: "- numeral m = - numeral n \ m = n" by simp lemma numeral_neq_neg_numeral: "numeral m \ - numeral n" by (simp add: eq_neg_iff_add_eq_0 numeral_plus_numeral) lemma neg_numeral_neq_numeral: "- numeral m \ numeral n" by (rule numeral_neq_neg_numeral [symmetric]) lemma zero_neq_neg_numeral: "0 \ - numeral n" by simp lemma neg_numeral_neq_zero: "- numeral n \ 0" by simp lemma one_neq_neg_numeral: "1 \ - numeral n" using numeral_neq_neg_numeral [of One n] by (simp add: numeral_One) lemma neg_numeral_neq_one: "- numeral n \ 1" using neg_numeral_neq_numeral [of n One] by (simp add: numeral_One) lemma neg_one_neq_numeral: "- 1 \ numeral n" using neg_numeral_neq_numeral [of One n] by (simp add: numeral_One) lemma numeral_neq_neg_one: "numeral n \ - 1" using numeral_neq_neg_numeral [of n One] by (simp add: numeral_One) lemma neg_one_eq_numeral_iff: "- 1 = - numeral n \ n = One" using neg_numeral_eq_iff [of One n] by (auto simp add: numeral_One) lemma numeral_eq_neg_one_iff: "- numeral n = - 1 \ n = One" using neg_numeral_eq_iff [of n One] by (auto simp add: numeral_One) lemma neg_one_neq_zero: "- 1 \ 0" by simp lemma zero_neq_neg_one: "0 \ - 1" by simp lemma neg_one_neq_one: "- 1 \ 1" using neg_numeral_neq_numeral [of One One] by (simp only: numeral_One not_False_eq_True) lemma one_neq_neg_one: "1 \ - 1" using numeral_neq_neg_numeral [of One One] by (simp only: numeral_One not_False_eq_True) lemmas eq_neg_numeral_simps [simp] = neg_numeral_eq_iff numeral_neq_neg_numeral neg_numeral_neq_numeral one_neq_neg_numeral neg_numeral_neq_one zero_neq_neg_numeral neg_numeral_neq_zero neg_one_neq_numeral numeral_neq_neg_one neg_one_eq_numeral_iff numeral_eq_neg_one_iff neg_one_neq_zero zero_neq_neg_one neg_one_neq_one one_neq_neg_one end subsubsection \Structures with negation and order: class \linordered_idom\\ context linordered_idom begin subclass ring_char_0 .. lemma neg_numeral_le_iff: "- numeral m \ - numeral n \ n \ m" by (simp only: neg_le_iff_le numeral_le_iff) lemma neg_numeral_less_iff: "- numeral m < - numeral n \ n < m" by (simp only: neg_less_iff_less numeral_less_iff) lemma neg_numeral_less_zero: "- numeral n < 0" by (simp only: neg_less_0_iff_less zero_less_numeral) lemma neg_numeral_le_zero: "- numeral n \ 0" by (simp only: neg_le_0_iff_le zero_le_numeral) lemma not_zero_less_neg_numeral: "\ 0 < - numeral n" by (simp only: not_less neg_numeral_le_zero) lemma not_zero_le_neg_numeral: "\ 0 \ - numeral n" by (simp only: not_le neg_numeral_less_zero) lemma neg_numeral_less_numeral: "- numeral m < numeral n" using neg_numeral_less_zero zero_less_numeral by (rule less_trans) lemma neg_numeral_le_numeral: "- numeral m \ numeral n" by (simp only: less_imp_le neg_numeral_less_numeral) lemma not_numeral_less_neg_numeral: "\ numeral m < - numeral n" by (simp only: not_less neg_numeral_le_numeral) lemma not_numeral_le_neg_numeral: "\ numeral m \ - numeral n" by (simp only: not_le neg_numeral_less_numeral) lemma neg_numeral_less_one: "- numeral m < 1" by (rule neg_numeral_less_numeral [of m One, unfolded numeral_One]) lemma neg_numeral_le_one: "- numeral m \ 1" by (rule neg_numeral_le_numeral [of m One, unfolded numeral_One]) lemma not_one_less_neg_numeral: "\ 1 < - numeral m" by (simp only: not_less neg_numeral_le_one) lemma not_one_le_neg_numeral: "\ 1 \ - numeral m" by (simp only: not_le neg_numeral_less_one) lemma not_numeral_less_neg_one: "\ numeral m < - 1" using not_numeral_less_neg_numeral [of m One] by (simp add: numeral_One) lemma not_numeral_le_neg_one: "\ numeral m \ - 1" using not_numeral_le_neg_numeral [of m One] by (simp add: numeral_One) lemma neg_one_less_numeral: "- 1 < numeral m" using neg_numeral_less_numeral [of One m] by (simp add: numeral_One) lemma neg_one_le_numeral: "- 1 \ numeral m" using neg_numeral_le_numeral [of One m] by (simp add: numeral_One) lemma neg_numeral_less_neg_one_iff: "- numeral m < - 1 \ m \ One" by (cases m) simp_all lemma neg_numeral_le_neg_one: "- numeral m \ - 1" by simp lemma not_neg_one_less_neg_numeral: "\ - 1 < - numeral m" by simp lemma not_neg_one_le_neg_numeral_iff: "\ - 1 \ - numeral m \ m \ One" by (cases m) simp_all lemma sub_non_negative: "sub n m \ 0 \ n \ m" by (simp only: sub_def le_diff_eq) simp lemma sub_positive: "sub n m > 0 \ n > m" by (simp only: sub_def less_diff_eq) simp lemma sub_non_positive: "sub n m \ 0 \ n \ m" by (simp only: sub_def diff_le_eq) simp lemma sub_negative: "sub n m < 0 \ n < m" by (simp only: sub_def diff_less_eq) simp lemmas le_neg_numeral_simps [simp] = neg_numeral_le_iff neg_numeral_le_numeral not_numeral_le_neg_numeral neg_numeral_le_zero not_zero_le_neg_numeral neg_numeral_le_one not_one_le_neg_numeral neg_one_le_numeral not_numeral_le_neg_one neg_numeral_le_neg_one not_neg_one_le_neg_numeral_iff lemma le_minus_one_simps [simp]: "- 1 \ 0" "- 1 \ 1" "\ 0 \ - 1" "\ 1 \ - 1" by simp_all lemmas less_neg_numeral_simps [simp] = neg_numeral_less_iff neg_numeral_less_numeral not_numeral_less_neg_numeral neg_numeral_less_zero not_zero_less_neg_numeral neg_numeral_less_one not_one_less_neg_numeral neg_one_less_numeral not_numeral_less_neg_one neg_numeral_less_neg_one_iff not_neg_one_less_neg_numeral lemma less_minus_one_simps [simp]: "- 1 < 0" "- 1 < 1" "\ 0 < - 1" "\ 1 < - 1" by (simp_all add: less_le) lemma abs_numeral [simp]: "\numeral n\ = numeral n" by simp lemma abs_neg_numeral [simp]: "\- numeral n\ = numeral n" by (simp only: abs_minus_cancel abs_numeral) lemma abs_neg_one [simp]: "\- 1\ = 1" by simp end subsubsection \Natural numbers\ lemma numeral_num_of_nat: "numeral (num_of_nat n) = n" if "n > 0" using that nat_of_num_numeral num_of_nat_inverse by simp lemma Suc_1 [simp]: "Suc 1 = 2" unfolding Suc_eq_plus1 by (rule one_add_one) lemma Suc_numeral [simp]: "Suc (numeral n) = numeral (n + One)" unfolding Suc_eq_plus1 by (rule numeral_plus_one) definition pred_numeral :: "num \ nat" where "pred_numeral k = numeral k - 1" declare [[code drop: pred_numeral]] lemma numeral_eq_Suc: "numeral k = Suc (pred_numeral k)" by (simp add: pred_numeral_def) lemma eval_nat_numeral: "numeral One = Suc 0" "numeral (Bit0 n) = Suc (numeral (BitM n))" "numeral (Bit1 n) = Suc (numeral (Bit0 n))" by (simp_all add: numeral.simps BitM_plus_one) lemma pred_numeral_simps [simp]: "pred_numeral One = 0" "pred_numeral (Bit0 k) = numeral (BitM k)" "pred_numeral (Bit1 k) = numeral (Bit0 k)" by (simp_all only: pred_numeral_def eval_nat_numeral diff_Suc_Suc diff_0) lemma pred_numeral_inc [simp]: "pred_numeral (Num.inc k) = numeral k" by (simp only: pred_numeral_def numeral_inc diff_add_inverse2) lemma numeral_2_eq_2: "2 = Suc (Suc 0)" by (simp add: eval_nat_numeral) lemma numeral_3_eq_3: "3 = Suc (Suc (Suc 0))" by (simp add: eval_nat_numeral) lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0" by (simp only: numeral_One One_nat_def) lemma Suc_nat_number_of_add: "Suc (numeral v + n) = numeral (v + One) + n" by simp lemma numerals: "Numeral1 = (1::nat)" "2 = Suc (Suc 0)" by (rule numeral_One) (rule numeral_2_eq_2) lemmas numeral_nat = eval_nat_numeral BitM.simps One_nat_def text \Comparisons involving \<^term>\Suc\.\ lemma eq_numeral_Suc [simp]: "numeral k = Suc n \ pred_numeral k = n" by (simp add: numeral_eq_Suc) lemma Suc_eq_numeral [simp]: "Suc n = numeral k \ n = pred_numeral k" by (simp add: numeral_eq_Suc) lemma less_numeral_Suc [simp]: "numeral k < Suc n \ pred_numeral k < n" by (simp add: numeral_eq_Suc) lemma less_Suc_numeral [simp]: "Suc n < numeral k \ n < pred_numeral k" by (simp add: numeral_eq_Suc) lemma le_numeral_Suc [simp]: "numeral k \ Suc n \ pred_numeral k \ n" by (simp add: numeral_eq_Suc) lemma le_Suc_numeral [simp]: "Suc n \ numeral k \ n \ pred_numeral k" by (simp add: numeral_eq_Suc) lemma diff_Suc_numeral [simp]: "Suc n - numeral k = n - pred_numeral k" by (simp add: numeral_eq_Suc) lemma diff_numeral_Suc [simp]: "numeral k - Suc n = pred_numeral k - n" by (simp add: numeral_eq_Suc) lemma max_Suc_numeral [simp]: "max (Suc n) (numeral k) = Suc (max n (pred_numeral k))" by (simp add: numeral_eq_Suc) lemma max_numeral_Suc [simp]: "max (numeral k) (Suc n) = Suc (max (pred_numeral k) n)" by (simp add: numeral_eq_Suc) lemma min_Suc_numeral [simp]: "min (Suc n) (numeral k) = Suc (min n (pred_numeral k))" by (simp add: numeral_eq_Suc) lemma min_numeral_Suc [simp]: "min (numeral k) (Suc n) = Suc (min (pred_numeral k) n)" by (simp add: numeral_eq_Suc) text \For \<^term>\case_nat\ and \<^term>\rec_nat\.\ lemma case_nat_numeral [simp]: "case_nat a f (numeral v) = (let pv = pred_numeral v in f pv)" by (simp add: numeral_eq_Suc) lemma case_nat_add_eq_if [simp]: "case_nat a f ((numeral v) + n) = (let pv = pred_numeral v in f (pv + n))" by (simp add: numeral_eq_Suc) lemma rec_nat_numeral [simp]: "rec_nat a f (numeral v) = (let pv = pred_numeral v in f pv (rec_nat a f pv))" by (simp add: numeral_eq_Suc Let_def) lemma rec_nat_add_eq_if [simp]: "rec_nat a f (numeral v + n) = (let pv = pred_numeral v in f (pv + n) (rec_nat a f (pv + n)))" by (simp add: numeral_eq_Suc Let_def) text \Case analysis on \<^term>\n < 2\.\ lemma less_2_cases: "n < 2 \ n = 0 \ n = Suc 0" by (auto simp add: numeral_2_eq_2) text \Removal of Small Numerals: 0, 1 and (in additive positions) 2.\ text \bh: Are these rules really a good idea?\ lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)" by simp lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)" by simp text \Can be used to eliminate long strings of Sucs, but not by default.\ lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" by simp lemmas nat_1_add_1 = one_add_one [where 'a=nat] (* legacy *) subsection \Particular lemmas concerning \<^term>\2\\ context linordered_field begin subclass field_char_0 .. lemma half_gt_zero_iff: "0 < a / 2 \ 0 < a" by (auto simp add: field_simps) lemma half_gt_zero [simp]: "0 < a \ 0 < a / 2" by (simp add: half_gt_zero_iff) end subsection \Numeral equations as default simplification rules\ declare (in numeral) numeral_One [simp] declare (in numeral) numeral_plus_numeral [simp] declare (in numeral) add_numeral_special [simp] declare (in neg_numeral) add_neg_numeral_simps [simp] declare (in neg_numeral) add_neg_numeral_special [simp] declare (in neg_numeral) diff_numeral_simps [simp] declare (in neg_numeral) diff_numeral_special [simp] declare (in semiring_numeral) numeral_times_numeral [simp] declare (in ring_1) mult_neg_numeral_simps [simp] subsubsection \Special Simplification for Constants\ text \These distributive laws move literals inside sums and differences.\ lemmas distrib_right_numeral [simp] = distrib_right [of _ _ "numeral v"] for v lemmas distrib_left_numeral [simp] = distrib_left [of "numeral v"] for v lemmas left_diff_distrib_numeral [simp] = left_diff_distrib [of _ _ "numeral v"] for v lemmas right_diff_distrib_numeral [simp] = right_diff_distrib [of "numeral v"] for v text \These are actually for fields, like real\ lemmas zero_less_divide_iff_numeral [simp, no_atp] = zero_less_divide_iff [of "numeral w"] for w lemmas divide_less_0_iff_numeral [simp, no_atp] = divide_less_0_iff [of "numeral w"] for w lemmas zero_le_divide_iff_numeral [simp, no_atp] = zero_le_divide_iff [of "numeral w"] for w lemmas divide_le_0_iff_numeral [simp, no_atp] = divide_le_0_iff [of "numeral w"] for w text \Replaces \inverse #nn\ by \1/#nn\. It looks strange, but then other simprocs simplify the quotient.\ lemmas inverse_eq_divide_numeral [simp] = inverse_eq_divide [of "numeral w"] for w lemmas inverse_eq_divide_neg_numeral [simp] = inverse_eq_divide [of "- numeral w"] for w text \These laws simplify inequalities, moving unary minus from a term into the literal.\ lemmas equation_minus_iff_numeral [no_atp] = equation_minus_iff [of "numeral v"] for v lemmas minus_equation_iff_numeral [no_atp] = minus_equation_iff [of _ "numeral v"] for v lemmas le_minus_iff_numeral [no_atp] = le_minus_iff [of "numeral v"] for v lemmas minus_le_iff_numeral [no_atp] = minus_le_iff [of _ "numeral v"] for v lemmas less_minus_iff_numeral [no_atp] = less_minus_iff [of "numeral v"] for v lemmas minus_less_iff_numeral [no_atp] = minus_less_iff [of _ "numeral v"] for v (* FIXME maybe simproc *) text \Cancellation of constant factors in comparisons (\<\ and \\\)\ lemmas mult_less_cancel_left_numeral [simp, no_atp] = mult_less_cancel_left [of "numeral v"] for v lemmas mult_less_cancel_right_numeral [simp, no_atp] = mult_less_cancel_right [of _ "numeral v"] for v lemmas mult_le_cancel_left_numeral [simp, no_atp] = mult_le_cancel_left [of "numeral v"] for v lemmas mult_le_cancel_right_numeral [simp, no_atp] = mult_le_cancel_right [of _ "numeral v"] for v text \Multiplying out constant divisors in comparisons (\<\, \\\ and \=\)\ named_theorems divide_const_simps "simplification rules to simplify comparisons involving constant divisors" lemmas le_divide_eq_numeral1 [simp,divide_const_simps] = pos_le_divide_eq [of "numeral w", OF zero_less_numeral] neg_le_divide_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas divide_le_eq_numeral1 [simp,divide_const_simps] = pos_divide_le_eq [of "numeral w", OF zero_less_numeral] neg_divide_le_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas less_divide_eq_numeral1 [simp,divide_const_simps] = pos_less_divide_eq [of "numeral w", OF zero_less_numeral] neg_less_divide_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas divide_less_eq_numeral1 [simp,divide_const_simps] = pos_divide_less_eq [of "numeral w", OF zero_less_numeral] neg_divide_less_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas eq_divide_eq_numeral1 [simp,divide_const_simps] = eq_divide_eq [of _ _ "numeral w"] eq_divide_eq [of _ _ "- numeral w"] for w lemmas divide_eq_eq_numeral1 [simp,divide_const_simps] = divide_eq_eq [of _ "numeral w"] divide_eq_eq [of _ "- numeral w"] for w subsubsection \Optional Simplification Rules Involving Constants\ text \Simplify quotients that are compared with a literal constant.\ lemmas le_divide_eq_numeral [divide_const_simps] = le_divide_eq [of "numeral w"] le_divide_eq [of "- numeral w"] for w lemmas divide_le_eq_numeral [divide_const_simps] = divide_le_eq [of _ _ "numeral w"] divide_le_eq [of _ _ "- numeral w"] for w lemmas less_divide_eq_numeral [divide_const_simps] = less_divide_eq [of "numeral w"] less_divide_eq [of "- numeral w"] for w lemmas divide_less_eq_numeral [divide_const_simps] = divide_less_eq [of _ _ "numeral w"] divide_less_eq [of _ _ "- numeral w"] for w lemmas eq_divide_eq_numeral [divide_const_simps] = eq_divide_eq [of "numeral w"] eq_divide_eq [of "- numeral w"] for w lemmas divide_eq_eq_numeral [divide_const_simps] = divide_eq_eq [of _ _ "numeral w"] divide_eq_eq [of _ _ "- numeral w"] for w text \Not good as automatic simprules because they cause case splits.\ lemmas [divide_const_simps] = le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1 subsection \Setting up simprocs\ lemma mult_numeral_1: "Numeral1 * a = a" for a :: "'a::semiring_numeral" by simp lemma mult_numeral_1_right: "a * Numeral1 = a" for a :: "'a::semiring_numeral" by simp lemma divide_numeral_1: "a / Numeral1 = a" for a :: "'a::field" by simp lemma inverse_numeral_1: "inverse Numeral1 = (Numeral1::'a::division_ring)" by simp text \ Theorem lists for the cancellation simprocs. The use of a binary numeral for 1 reduces the number of special cases. \ lemma mult_1s_semiring_numeral: "Numeral1 * a = a" "a * Numeral1 = a" for a :: "'a::semiring_numeral" by simp_all lemma mult_1s_ring_1: "- Numeral1 * b = - b" "b * - Numeral1 = - b" for b :: "'a::ring_1" by simp_all lemmas mult_1s = mult_1s_semiring_numeral mult_1s_ring_1 setup \ Reorient_Proc.add (fn Const (\<^const_name>\numeral\, _) $ _ => true | Const (\<^const_name>\uminus\, _) $ (Const (\<^const_name>\numeral\, _) $ _) => true | _ => false) \ simproc_setup reorient_numeral ("numeral w = x" | "- numeral w = y") = Reorient_Proc.proc subsubsection \Simplification of arithmetic operations on integer constants\ lemmas arith_special = (* already declared simp above *) add_numeral_special add_neg_numeral_special diff_numeral_special lemmas arith_extra_simps = (* rules already in simpset *) numeral_plus_numeral add_neg_numeral_simps add_0_left add_0_right minus_zero diff_numeral_simps diff_0 diff_0_right numeral_times_numeral mult_neg_numeral_simps mult_zero_left mult_zero_right abs_numeral abs_neg_numeral text \ For making a minimal simpset, one must include these default simprules. Also include \simp_thms\. \ lemmas arith_simps = add_num_simps mult_num_simps sub_num_simps BitM.simps dbl_simps dbl_inc_simps dbl_dec_simps abs_zero abs_one arith_extra_simps lemmas more_arith_simps = neg_le_iff_le minus_zero left_minus right_minus mult_1_left mult_1_right mult_minus_left mult_minus_right minus_add_distrib minus_minus mult.assoc lemmas of_nat_simps = of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult text \Simplification of relational operations.\ lemmas eq_numeral_extra = zero_neq_one one_neq_zero lemmas rel_simps = le_num_simps less_num_simps eq_num_simps le_numeral_simps le_neg_numeral_simps le_minus_one_simps le_numeral_extra less_numeral_simps less_neg_numeral_simps less_minus_one_simps less_numeral_extra eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)" \ \Unfold all \let\s involving constants\ unfolding Let_def .. lemma Let_neg_numeral [simp]: "Let (- numeral v) f = f (- numeral v)" \ \Unfold all \let\s involving constants\ unfolding Let_def .. declaration \ let fun number_of ctxt T n = if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\numeral\)) then raise CTERM ("number_of", []) else Numeral.mk_cnumber (Thm.ctyp_of ctxt T) n; in K ( Lin_Arith.set_number_of number_of #> Lin_Arith.add_simps @{thms arith_simps more_arith_simps rel_simps pred_numeral_simps arith_special numeral_One of_nat_simps uminus_numeral_One Suc_numeral Let_numeral Let_neg_numeral Let_0 Let_1 le_Suc_numeral le_numeral_Suc less_Suc_numeral less_numeral_Suc Suc_eq_numeral eq_numeral_Suc mult_Suc mult_Suc_right of_nat_numeral}) end \ subsubsection \Simplification of arithmetic when nested to the right\ lemma add_numeral_left [simp]: "numeral v + (numeral w + z) = (numeral(v + w) + z)" by (simp_all add: add.assoc [symmetric]) lemma add_neg_numeral_left [simp]: "numeral v + (- numeral w + y) = (sub v w + y)" "- numeral v + (numeral w + y) = (sub w v + y)" "- numeral v + (- numeral w + y) = (- numeral(v + w) + y)" by (simp_all add: add.assoc [symmetric]) lemma mult_numeral_left_semiring_numeral: "numeral v * (numeral w * z) = (numeral(v * w) * z :: 'a::semiring_numeral)" by (simp add: mult.assoc [symmetric]) lemma mult_numeral_left_ring_1: "- numeral v * (numeral w * y) = (- numeral(v * w) * y :: 'a::ring_1)" "numeral v * (- numeral w * y) = (- numeral(v * w) * y :: 'a::ring_1)" "- numeral v * (- numeral w * y) = (numeral(v * w) * y :: 'a::ring_1)" by (simp_all add: mult.assoc [symmetric]) lemmas mult_numeral_left [simp] = mult_numeral_left_semiring_numeral mult_numeral_left_ring_1 hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec subsection \Code module namespace\ code_identifier code_module Num \ (SML) Arith and (OCaml) Arith and (Haskell) Arith subsection \Printing of evaluated natural numbers as numerals\ lemma [code_post]: "Suc 0 = 1" "Suc 1 = 2" "Suc (numeral n) = numeral (Num.inc n)" by (simp_all add: numeral_inc) lemmas [code_post] = Num.inc.simps end diff --git a/src/HOL/Transfer.thy b/src/HOL/Transfer.thy --- a/src/HOL/Transfer.thy +++ b/src/HOL/Transfer.thy @@ -1,649 +1,654 @@ (* Title: HOL/Transfer.thy Author: Brian Huffman, TU Muenchen Author: Ondrej Kuncar, TU Muenchen *) section \Generic theorem transfer using relations\ theory Transfer imports Basic_BNF_LFPs Hilbert_Choice Metis begin subsection \Relator for function space\ bundle lifting_syntax begin notation rel_fun (infixr "===>" 55) notation map_fun (infixr "--->" 55) end context includes lifting_syntax begin lemma rel_funD2: assumes "rel_fun A B f g" and "A x x" shows "B (f x) (g x)" using assms by (rule rel_funD) lemma rel_funE: assumes "rel_fun A B f g" and "A x y" obtains "B (f x) (g y)" using assms by (simp add: rel_fun_def) lemmas rel_fun_eq = fun.rel_eq lemma rel_fun_eq_rel: shows "rel_fun (=) R = (\f g. \x. R (f x) (g x))" by (simp add: rel_fun_def) subsection \Transfer method\ text \Explicit tag for relation membership allows for backward proof methods.\ definition Rel :: "('a \ 'b \ bool) \ 'a \ 'b \ bool" where "Rel r \ r" text \Handling of equality relations\ definition is_equality :: "('a \ 'a \ bool) \ bool" where "is_equality R \ R = (=)" lemma is_equality_eq: "is_equality (=)" unfolding is_equality_def by simp text \Reverse implication for monotonicity rules\ definition rev_implies where "rev_implies x y \ (y \ x)" text \Handling of meta-logic connectives\ definition transfer_forall where "transfer_forall \ All" definition transfer_implies where "transfer_implies \ (\)" definition transfer_bforall :: "('a \ bool) \ ('a \ bool) \ bool" where "transfer_bforall \ (\P Q. \x. P x \ Q x)" lemma transfer_forall_eq: "(\x. P x) \ Trueprop (transfer_forall (\x. P x))" unfolding atomize_all transfer_forall_def .. lemma transfer_implies_eq: "(A \ B) \ Trueprop (transfer_implies A B)" unfolding atomize_imp transfer_implies_def .. lemma transfer_bforall_unfold: "Trueprop (transfer_bforall P (\x. Q x)) \ (\x. P x \ Q x)" unfolding transfer_bforall_def atomize_imp atomize_all .. lemma transfer_start: "\P; Rel (=) P Q\ \ Q" unfolding Rel_def by simp lemma transfer_start': "\P; Rel (\) P Q\ \ Q" unfolding Rel_def by simp lemma transfer_prover_start: "\x = x'; Rel R x' y\ \ Rel R x y" by simp lemma untransfer_start: "\Q; Rel (=) P Q\ \ P" unfolding Rel_def by simp lemma Rel_eq_refl: "Rel (=) x x" unfolding Rel_def .. lemma Rel_app: assumes "Rel (A ===> B) f g" and "Rel A x y" shows "Rel B (f x) (g y)" using assms unfolding Rel_def rel_fun_def by fast lemma Rel_abs: assumes "\x y. Rel A x y \ Rel B (f x) (g y)" shows "Rel (A ===> B) (\x. f x) (\y. g y)" using assms unfolding Rel_def rel_fun_def by fast subsection \Predicates on relations, i.e. ``class constraints''\ definition left_total :: "('a \ 'b \ bool) \ bool" where "left_total R \ (\x. \y. R x y)" definition left_unique :: "('a \ 'b \ bool) \ bool" where "left_unique R \ (\x y z. R x z \ R y z \ x = y)" definition right_total :: "('a \ 'b \ bool) \ bool" where "right_total R \ (\y. \x. R x y)" definition right_unique :: "('a \ 'b \ bool) \ bool" where "right_unique R \ (\x y z. R x y \ R x z \ y = z)" definition bi_total :: "('a \ 'b \ bool) \ bool" where "bi_total R \ (\x. \y. R x y) \ (\y. \x. R x y)" definition bi_unique :: "('a \ 'b \ bool) \ bool" where "bi_unique R \ (\x y z. R x y \ R x z \ y = z) \ (\x y z. R x z \ R y z \ x = y)" lemma left_uniqueI: "(\x y z. \ A x z; A y z \ \ x = y) \ left_unique A" unfolding left_unique_def by blast lemma left_uniqueD: "\ left_unique A; A x z; A y z \ \ x = y" unfolding left_unique_def by blast lemma left_totalI: "(\x. \y. R x y) \ left_total R" unfolding left_total_def by blast lemma left_totalE: assumes "left_total R" obtains "(\x. \y. R x y)" using assms unfolding left_total_def by blast lemma bi_uniqueDr: "\ bi_unique A; A x y; A x z \ \ y = z" by(simp add: bi_unique_def) lemma bi_uniqueDl: "\ bi_unique A; A x y; A z y \ \ x = z" by(simp add: bi_unique_def) lemma right_uniqueI: "(\x y z. \ A x y; A x z \ \ y = z) \ right_unique A" unfolding right_unique_def by fast lemma right_uniqueD: "\ right_unique A; A x y; A x z \ \ y = z" unfolding right_unique_def by fast lemma right_totalI: "(\y. \x. A x y) \ right_total A" by(simp add: right_total_def) lemma right_totalE: assumes "right_total A" obtains x where "A x y" using assms by(auto simp add: right_total_def) lemma right_total_alt_def2: "right_total R \ ((R ===> (\)) ===> (\)) All All" unfolding right_total_def rel_fun_def apply (rule iffI, fast) apply (rule allI) apply (drule_tac x="\x. True" in spec) apply (drule_tac x="\y. \x. R x y" in spec) apply fast done lemma right_unique_alt_def2: "right_unique R \ (R ===> R ===> (\)) (=) (=)" unfolding right_unique_def rel_fun_def by auto lemma bi_total_alt_def2: "bi_total R \ ((R ===> (=)) ===> (=)) All All" unfolding bi_total_def rel_fun_def apply (rule iffI, fast) apply safe apply (drule_tac x="\x. \y. R x y" in spec) apply (drule_tac x="\y. True" in spec) apply fast apply (drule_tac x="\x. True" in spec) apply (drule_tac x="\y. \x. R x y" in spec) apply fast done lemma bi_unique_alt_def2: "bi_unique R \ (R ===> R ===> (=)) (=) (=)" unfolding bi_unique_def rel_fun_def by auto lemma [simp]: shows left_unique_conversep: "left_unique A\\ \ right_unique A" and right_unique_conversep: "right_unique A\\ \ left_unique A" by(auto simp add: left_unique_def right_unique_def) lemma [simp]: shows left_total_conversep: "left_total A\\ \ right_total A" and right_total_conversep: "right_total A\\ \ left_total A" by(simp_all add: left_total_def right_total_def) lemma bi_unique_conversep [simp]: "bi_unique R\\ = bi_unique R" by(auto simp add: bi_unique_def) lemma bi_total_conversep [simp]: "bi_total R\\ = bi_total R" by(auto simp add: bi_total_def) lemma right_unique_alt_def: "right_unique R = (conversep R OO R \ (=))" unfolding right_unique_def by blast lemma left_unique_alt_def: "left_unique R = (R OO (conversep R) \ (=))" unfolding left_unique_def by blast lemma right_total_alt_def: "right_total R = (conversep R OO R \ (=))" unfolding right_total_def by blast lemma left_total_alt_def: "left_total R = (R OO conversep R \ (=))" unfolding left_total_def by blast lemma bi_total_alt_def: "bi_total A = (left_total A \ right_total A)" unfolding left_total_def right_total_def bi_total_def by blast lemma bi_unique_alt_def: "bi_unique A = (left_unique A \ right_unique A)" unfolding left_unique_def right_unique_def bi_unique_def by blast lemma bi_totalI: "left_total R \ right_total R \ bi_total R" unfolding bi_total_alt_def .. lemma bi_uniqueI: "left_unique R \ right_unique R \ bi_unique R" unfolding bi_unique_alt_def .. end lemma is_equality_lemma: "(\R. is_equality R \ PROP (P R)) \ PROP (P (=))" apply (unfold is_equality_def) apply (rule equal_intr_rule) apply (drule meta_spec) apply (erule meta_mp) apply (rule refl) apply simp done lemma Domainp_lemma: "(\R. Domainp T = R \ PROP (P R)) \ PROP (P (Domainp T))" apply (rule equal_intr_rule) apply (drule meta_spec) apply (erule meta_mp) apply (rule refl) apply simp done ML_file \Tools/Transfer/transfer.ML\ declare refl [transfer_rule] hide_const (open) Rel context includes lifting_syntax begin text \Handling of domains\ lemma Domainp_iff: "Domainp T x \ (\y. T x y)" by auto lemma Domainp_refl[transfer_domain_rule]: "Domainp T = Domainp T" .. lemma Domain_eq_top[transfer_domain_rule]: "Domainp (=) = top" by auto lemma Domainp_pred_fun_eq[relator_domain]: assumes "left_unique T" shows "Domainp (T ===> S) = pred_fun (Domainp T) (Domainp S)" using assms unfolding rel_fun_def Domainp_iff[abs_def] left_unique_def fun_eq_iff pred_fun_def apply safe apply blast apply (subst all_comm) apply (rule choice) apply blast done text \Properties are preserved by relation composition.\ lemma OO_def: "R OO S = (\x z. \y. R x y \ S y z)" by auto lemma bi_total_OO: "\bi_total A; bi_total B\ \ bi_total (A OO B)" unfolding bi_total_def OO_def by fast lemma bi_unique_OO: "\bi_unique A; bi_unique B\ \ bi_unique (A OO B)" unfolding bi_unique_def OO_def by blast lemma right_total_OO: "\right_total A; right_total B\ \ right_total (A OO B)" unfolding right_total_def OO_def by fast lemma right_unique_OO: "\right_unique A; right_unique B\ \ right_unique (A OO B)" unfolding right_unique_def OO_def by fast lemma left_total_OO: "left_total R \ left_total S \ left_total (R OO S)" unfolding left_total_def OO_def by fast lemma left_unique_OO: "left_unique R \ left_unique S \ left_unique (R OO S)" unfolding left_unique_def OO_def by blast subsection \Properties of relators\ lemma left_total_eq[transfer_rule]: "left_total (=)" unfolding left_total_def by blast lemma left_unique_eq[transfer_rule]: "left_unique (=)" unfolding left_unique_def by blast lemma right_total_eq [transfer_rule]: "right_total (=)" unfolding right_total_def by simp lemma right_unique_eq [transfer_rule]: "right_unique (=)" unfolding right_unique_def by simp lemma bi_total_eq[transfer_rule]: "bi_total (=)" unfolding bi_total_def by simp lemma bi_unique_eq[transfer_rule]: "bi_unique (=)" unfolding bi_unique_def by simp lemma left_total_fun[transfer_rule]: "\left_unique A; left_total B\ \ left_total (A ===> B)" unfolding left_total_def rel_fun_def apply (rule allI, rename_tac f) apply (rule_tac x="\y. SOME z. B (f (THE x. A x y)) z" in exI) apply clarify apply (subgoal_tac "(THE x. A x y) = x", simp) apply (rule someI_ex) apply (simp) apply (rule the_equality) apply assumption apply (simp add: left_unique_def) done lemma left_unique_fun[transfer_rule]: "\left_total A; left_unique B\ \ left_unique (A ===> B)" unfolding left_total_def left_unique_def rel_fun_def by (clarify, rule ext, fast) lemma right_total_fun [transfer_rule]: "\right_unique A; right_total B\ \ right_total (A ===> B)" unfolding right_total_def rel_fun_def apply (rule allI, rename_tac g) apply (rule_tac x="\x. SOME z. B z (g (THE y. A x y))" in exI) apply clarify apply (subgoal_tac "(THE y. A x y) = y", simp) apply (rule someI_ex) apply (simp) apply (rule the_equality) apply assumption apply (simp add: right_unique_def) done lemma right_unique_fun [transfer_rule]: "\right_total A; right_unique B\ \ right_unique (A ===> B)" unfolding right_total_def right_unique_def rel_fun_def by (clarify, rule ext, fast) lemma bi_total_fun[transfer_rule]: "\bi_unique A; bi_total B\ \ bi_total (A ===> B)" unfolding bi_unique_alt_def bi_total_alt_def by (blast intro: right_total_fun left_total_fun) lemma bi_unique_fun[transfer_rule]: "\bi_total A; bi_unique B\ \ bi_unique (A ===> B)" unfolding bi_unique_alt_def bi_total_alt_def by (blast intro: right_unique_fun left_unique_fun) end lemma if_conn: "(if P \ Q then t else e) = (if P then if Q then t else e else e)" "(if P \ Q then t else e) = (if P then t else if Q then t else e)" "(if P \ Q then t else e) = (if P then if Q then t else e else t)" "(if \ P then t else e) = (if P then e else t)" by auto ML_file \Tools/Transfer/transfer_bnf.ML\ ML_file \Tools/BNF/bnf_fp_rec_sugar_transfer.ML\ declare pred_fun_def [simp] declare rel_fun_eq [relator_eq] (* Delete the automated generated rule from the bnf command; we have a more general rule (Domainp_pred_fun_eq) that subsumes it. *) declare fun.Domainp_rel[relator_domain del] subsection \Transfer rules\ context includes lifting_syntax begin lemma Domainp_forall_transfer [transfer_rule]: assumes "right_total A" shows "((A ===> (=)) ===> (=)) (transfer_bforall (Domainp A)) transfer_forall" using assms unfolding right_total_def unfolding transfer_forall_def transfer_bforall_def rel_fun_def Domainp_iff by fast text \Transfer rules using implication instead of equality on booleans.\ lemma transfer_forall_transfer [transfer_rule]: "bi_total A \ ((A ===> (=)) ===> (=)) transfer_forall transfer_forall" "right_total A \ ((A ===> (=)) ===> implies) transfer_forall transfer_forall" "right_total A \ ((A ===> implies) ===> implies) transfer_forall transfer_forall" "bi_total A \ ((A ===> (=)) ===> rev_implies) transfer_forall transfer_forall" "bi_total A \ ((A ===> rev_implies) ===> rev_implies) transfer_forall transfer_forall" unfolding transfer_forall_def rev_implies_def rel_fun_def right_total_def bi_total_def by fast+ lemma transfer_implies_transfer [transfer_rule]: "((=) ===> (=) ===> (=) ) transfer_implies transfer_implies" "(rev_implies ===> implies ===> implies ) transfer_implies transfer_implies" "(rev_implies ===> (=) ===> implies ) transfer_implies transfer_implies" "((=) ===> implies ===> implies ) transfer_implies transfer_implies" "((=) ===> (=) ===> implies ) transfer_implies transfer_implies" "(implies ===> rev_implies ===> rev_implies) transfer_implies transfer_implies" "(implies ===> (=) ===> rev_implies) transfer_implies transfer_implies" "((=) ===> rev_implies ===> rev_implies) transfer_implies transfer_implies" "((=) ===> (=) ===> rev_implies) transfer_implies transfer_implies" unfolding transfer_implies_def rev_implies_def rel_fun_def by auto lemma eq_imp_transfer [transfer_rule]: "right_unique A \ (A ===> A ===> (\)) (=) (=)" unfolding right_unique_alt_def2 . text \Transfer rules using equality.\ lemma left_unique_transfer [transfer_rule]: assumes "right_total A" assumes "right_total B" assumes "bi_unique A" shows "((A ===> B ===> (=)) ===> implies) left_unique left_unique" using assms unfolding left_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def by metis lemma eq_transfer [transfer_rule]: assumes "bi_unique A" shows "(A ===> A ===> (=)) (=) (=)" using assms unfolding bi_unique_def rel_fun_def by auto lemma right_total_Ex_transfer[transfer_rule]: assumes "right_total A" shows "((A ===> (=)) ===> (=)) (Bex (Collect (Domainp A))) Ex" using assms unfolding right_total_def Bex_def rel_fun_def Domainp_iff[abs_def] by fast lemma right_total_All_transfer[transfer_rule]: assumes "right_total A" shows "((A ===> (=)) ===> (=)) (Ball (Collect (Domainp A))) All" using assms unfolding right_total_def Ball_def rel_fun_def Domainp_iff[abs_def] by fast +context + includes lifting_syntax +begin + lemma right_total_fun_eq_transfer: - includes lifting_syntax assumes [transfer_rule]: "right_total A" "bi_unique B" shows "((A ===> B) ===> (A ===> B) ===> (=)) (\f g. \x\Collect(Domainp A). f x = g x) (=)" unfolding fun_eq_iff by transfer_prover +end + lemma All_transfer [transfer_rule]: assumes "bi_total A" shows "((A ===> (=)) ===> (=)) All All" using assms unfolding bi_total_def rel_fun_def by fast lemma Ex_transfer [transfer_rule]: assumes "bi_total A" shows "((A ===> (=)) ===> (=)) Ex Ex" using assms unfolding bi_total_def rel_fun_def by fast lemma Ex1_parametric [transfer_rule]: assumes [transfer_rule]: "bi_unique A" "bi_total A" shows "((A ===> (=)) ===> (=)) Ex1 Ex1" unfolding Ex1_def[abs_def] by transfer_prover declare If_transfer [transfer_rule] lemma Let_transfer [transfer_rule]: "(A ===> (A ===> B) ===> B) Let Let" unfolding rel_fun_def by simp declare id_transfer [transfer_rule] declare comp_transfer [transfer_rule] lemma curry_transfer [transfer_rule]: "((rel_prod A B ===> C) ===> A ===> B ===> C) curry curry" unfolding curry_def by transfer_prover lemma fun_upd_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd" unfolding fun_upd_def [abs_def] by transfer_prover lemma case_nat_transfer [transfer_rule]: "(A ===> ((=) ===> A) ===> (=) ===> A) case_nat case_nat" unfolding rel_fun_def by (simp split: nat.split) lemma rec_nat_transfer [transfer_rule]: "(A ===> ((=) ===> A ===> A) ===> (=) ===> A) rec_nat rec_nat" unfolding rel_fun_def by (clarsimp, rename_tac n, induct_tac n, simp_all) lemma funpow_transfer [transfer_rule]: "((=) ===> (A ===> A) ===> (A ===> A)) compow compow" unfolding funpow_def by transfer_prover lemma mono_transfer[transfer_rule]: assumes [transfer_rule]: "bi_total A" assumes [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" assumes [transfer_rule]: "(B ===> B ===> (=)) (\) (\)" shows "((A ===> B) ===> (=)) mono mono" unfolding mono_def[abs_def] by transfer_prover lemma right_total_relcompp_transfer[transfer_rule]: assumes [transfer_rule]: "right_total B" shows "((A ===> B ===> (=)) ===> (B ===> C ===> (=)) ===> A ===> C ===> (=)) (\R S x z. \y\Collect (Domainp B). R x y \ S y z) (OO)" unfolding OO_def[abs_def] by transfer_prover lemma relcompp_transfer[transfer_rule]: assumes [transfer_rule]: "bi_total B" shows "((A ===> B ===> (=)) ===> (B ===> C ===> (=)) ===> A ===> C ===> (=)) (OO) (OO)" unfolding OO_def[abs_def] by transfer_prover lemma right_total_Domainp_transfer[transfer_rule]: assumes [transfer_rule]: "right_total B" shows "((A ===> B ===> (=)) ===> A ===> (=)) (\T x. \y\Collect(Domainp B). T x y) Domainp" apply(subst(2) Domainp_iff[abs_def]) by transfer_prover lemma Domainp_transfer[transfer_rule]: assumes [transfer_rule]: "bi_total B" shows "((A ===> B ===> (=)) ===> A ===> (=)) Domainp Domainp" unfolding Domainp_iff[abs_def] by transfer_prover lemma reflp_transfer[transfer_rule]: "bi_total A \ ((A ===> A ===> (=)) ===> (=)) reflp reflp" "right_total A \ ((A ===> A ===> implies) ===> implies) reflp reflp" "right_total A \ ((A ===> A ===> (=)) ===> implies) reflp reflp" "bi_total A \ ((A ===> A ===> rev_implies) ===> rev_implies) reflp reflp" "bi_total A \ ((A ===> A ===> (=)) ===> rev_implies) reflp reflp" unfolding reflp_def[abs_def] rev_implies_def bi_total_def right_total_def rel_fun_def by fast+ lemma right_unique_transfer [transfer_rule]: "\ right_total A; right_total B; bi_unique B \ \ ((A ===> B ===> (=)) ===> implies) right_unique right_unique" unfolding right_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def by metis lemma left_total_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" "bi_total B" shows "((A ===> B ===> (=)) ===> (=)) left_total left_total" unfolding left_total_def[abs_def] by transfer_prover lemma right_total_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" "bi_total B" shows "((A ===> B ===> (=)) ===> (=)) right_total right_total" unfolding right_total_def[abs_def] by transfer_prover lemma left_unique_parametric [transfer_rule]: assumes [transfer_rule]: "bi_unique A" "bi_total A" "bi_total B" shows "((A ===> B ===> (=)) ===> (=)) left_unique left_unique" unfolding left_unique_def[abs_def] by transfer_prover lemma prod_pred_parametric [transfer_rule]: "((A ===> (=)) ===> (B ===> (=)) ===> rel_prod A B ===> (=)) pred_prod pred_prod" unfolding prod.pred_set[abs_def] Basic_BNFs.fsts_def Basic_BNFs.snds_def fstsp.simps sndsp.simps by simp transfer_prover lemma apfst_parametric [transfer_rule]: "((A ===> B) ===> rel_prod A C ===> rel_prod B C) apfst apfst" unfolding apfst_def[abs_def] by transfer_prover lemma rel_fun_eq_eq_onp: "((=) ===> eq_onp P) = eq_onp (\f. \x. P(f x))" unfolding eq_onp_def rel_fun_def by auto lemma rel_fun_eq_onp_rel: shows "((eq_onp R) ===> S) = (\f g. \x. R x \ S (f x) (g x))" by (auto simp add: eq_onp_def rel_fun_def) lemma eq_onp_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "((A ===> (=)) ===> A ===> A ===> (=)) eq_onp eq_onp" unfolding eq_onp_def[abs_def] by transfer_prover lemma rtranclp_parametric [transfer_rule]: assumes "bi_unique A" "bi_total A" shows "((A ===> A ===> (=)) ===> A ===> A ===> (=)) rtranclp rtranclp" proof(rule rel_funI iffI)+ fix R :: "'a \ 'a \ bool" and R' x y x' y' assume R: "(A ===> A ===> (=)) R R'" and "A x x'" { assume "R\<^sup>*\<^sup>* x y" "A y y'" thus "R'\<^sup>*\<^sup>* x' y'" proof(induction arbitrary: y') case base with \bi_unique A\ \A x x'\ have "x' = y'" by(rule bi_uniqueDr) thus ?case by simp next case (step y z z') from \bi_total A\ obtain y' where "A y y'" unfolding bi_total_def by blast hence "R'\<^sup>*\<^sup>* x' y'" by(rule step.IH) moreover from R \A y y'\ \A z z'\ \R y z\ have "R' y' z'" by(auto dest: rel_funD) ultimately show ?case .. qed next assume "R'\<^sup>*\<^sup>* x' y'" "A y y'" thus "R\<^sup>*\<^sup>* x y" proof(induction arbitrary: y) case base with \bi_unique A\ \A x x'\ have "x = y" by(rule bi_uniqueDl) thus ?case by simp next case (step y' z' z) from \bi_total A\ obtain y where "A y y'" unfolding bi_total_def by blast hence "R\<^sup>*\<^sup>* x y" by(rule step.IH) moreover from R \A y y'\ \A z z'\ \R' y' z'\ have "R y z" by(auto dest: rel_funD) ultimately show ?case .. qed } qed lemma right_unique_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" "bi_unique B" "bi_total B" shows "((A ===> B ===> (=)) ===> (=)) right_unique right_unique" unfolding right_unique_def[abs_def] by transfer_prover lemma map_fun_parametric [transfer_rule]: "((A ===> B) ===> (C ===> D) ===> (B ===> C) ===> A ===> D) map_fun map_fun" unfolding map_fun_def[abs_def] by transfer_prover end subsection \\<^const>\of_nat\\ lemma transfer_rule_of_nat: fixes R :: "'a::semiring_1 \ 'b::semiring_1 \ bool" assumes [transfer_rule]: "R 0 0" "R 1 1" "rel_fun R (rel_fun R R) plus plus" shows "rel_fun HOL.eq R of_nat of_nat" by (unfold of_nat_def [abs_def]) transfer_prover end diff --git a/src/HOL/ex/Word_Type.thy b/src/HOL/ex/Word_Type.thy --- a/src/HOL/ex/Word_Type.thy +++ b/src/HOL/ex/Word_Type.thy @@ -1,598 +1,606 @@ (* Author: Florian Haftmann, TUM *) section \Proof of concept for algebraically founded bit word types\ theory Word_Type imports Main "HOL-ex.Bit_Lists" "HOL-Library.Type_Length" begin subsection \Preliminaries\ lemma take_bit_uminus: "take_bit n (- (take_bit n k)) = take_bit n (- k)" for k :: int by (simp add: take_bit_eq_mod mod_minus_eq) lemma take_bit_minus: "take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)" for k l :: int by (simp add: take_bit_eq_mod mod_diff_eq) lemma take_bit_nonnegative [simp]: "take_bit n k \ 0" for k :: int by (simp add: take_bit_eq_mod) definition signed_take_bit :: "nat \ int \ int" where signed_take_bit_eq_take_bit: "signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n" lemma signed_take_bit_eq_take_bit': "signed_take_bit (n - Suc 0) k = take_bit n (k + 2 ^ (n - 1)) - 2 ^ (n - 1)" if "n > 0" using that by (simp add: signed_take_bit_eq_take_bit) lemma signed_take_bit_0 [simp]: "signed_take_bit 0 k = - (k mod 2)" proof (cases "even k") case True then have "odd (k + 1)" by simp then have "(k + 1) mod 2 = 1" by (simp add: even_iff_mod_2_eq_zero) with True show ?thesis by (simp add: signed_take_bit_eq_take_bit) next case False then show ?thesis by (simp add: signed_take_bit_eq_take_bit odd_iff_mod_2_eq_one) qed lemma signed_take_bit_Suc [simp]: "signed_take_bit (Suc n) k = signed_take_bit n (k div 2) * 2 + k mod 2" by (simp add: odd_iff_mod_2_eq_one signed_take_bit_eq_take_bit algebra_simps) lemma signed_take_bit_of_0 [simp]: "signed_take_bit n 0 = 0" by (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod) lemma signed_take_bit_of_minus_1 [simp]: "signed_take_bit n (- 1) = - 1" by (induct n) simp_all lemma signed_take_bit_eq_iff_take_bit_eq: "signed_take_bit (n - Suc 0) k = signed_take_bit (n - Suc 0) l \ take_bit n k = take_bit n l" (is "?P \ ?Q") if "n > 0" proof - from that obtain m where m: "n = Suc m" by (cases n) auto show ?thesis proof assume ?Q have "take_bit (Suc m) (k + 2 ^ m) = take_bit (Suc m) (take_bit (Suc m) k + take_bit (Suc m) (2 ^ m))" by (simp only: take_bit_add) also have "\ = take_bit (Suc m) (take_bit (Suc m) l + take_bit (Suc m) (2 ^ m))" by (simp only: \?Q\ m [symmetric]) also have "\ = take_bit (Suc m) (l + 2 ^ m)" by (simp only: take_bit_add) finally show ?P by (simp only: signed_take_bit_eq_take_bit m) simp next assume ?P with that have "(k + 2 ^ (n - Suc 0)) mod 2 ^ n = (l + 2 ^ (n - Suc 0)) mod 2 ^ n" by (simp add: signed_take_bit_eq_take_bit' take_bit_eq_mod) then have "(i + (k + 2 ^ (n - Suc 0))) mod 2 ^ n = (i + (l + 2 ^ (n - Suc 0))) mod 2 ^ n" for i by (metis mod_add_eq) then have "k mod 2 ^ n = l mod 2 ^ n" by (metis add_diff_cancel_right' uminus_add_conv_diff) then show ?Q by (simp add: take_bit_eq_mod) qed qed subsection \Bit strings as quotient type\ subsubsection \Basic properties\ quotient_type (overloaded) 'a word = int / "\k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len0) l" by (auto intro!: equivpI reflpI sympI transpI) instantiation word :: (len0) "{semiring_numeral, comm_semiring_0, comm_ring}" begin lift_definition zero_word :: "'a word" is 0 . lift_definition one_word :: "'a word" is 1 . lift_definition plus_word :: "'a word \ 'a word \ 'a word" is plus by (subst take_bit_add [symmetric]) (simp add: take_bit_add) lift_definition uminus_word :: "'a word \ 'a word" is uminus by (subst take_bit_uminus [symmetric]) (simp add: take_bit_uminus) lift_definition minus_word :: "'a word \ 'a word \ 'a word" is minus by (subst take_bit_minus [symmetric]) (simp add: take_bit_minus) lift_definition times_word :: "'a word \ 'a word \ 'a word" is times by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) instance by standard (transfer; simp add: algebra_simps)+ end instance word :: (len) comm_ring_1 by standard (transfer; simp)+ quickcheck_generator word constructors: "zero_class.zero :: ('a::len0) word", "numeral :: num \ ('a::len0) word", "uminus :: ('a::len0) word \ ('a::len0) word" subsubsection \Conversions\ -lemma [transfer_rule]: - "rel_fun HOL.eq (pcr_word :: int \ 'a::len word \ bool) numeral numeral" -proof - - note transfer_rule_numeral [transfer_rule] - show ?thesis by transfer_prover -qed +context + includes lifting_syntax + notes transfer_rule_numeral [transfer_rule] + transfer_rule_of_nat [transfer_rule] + transfer_rule_of_int [transfer_rule] +begin lemma [transfer_rule]: - "rel_fun HOL.eq pcr_word int of_nat" -proof - - note transfer_rule_of_nat [transfer_rule] - show ?thesis by transfer_prover -qed + "((=) ===> (pcr_word :: int \ 'a::len word \ bool)) numeral numeral" + by transfer_prover + +lemma [transfer_rule]: + "((=) ===> pcr_word) int of_nat" + by transfer_prover lemma [transfer_rule]: - "rel_fun HOL.eq pcr_word (\k. k) of_int" + "((=) ===> pcr_word) (\k. k) of_int" proof - - note transfer_rule_of_int [transfer_rule] - have "rel_fun HOL.eq pcr_word (of_int :: int \ int) (of_int :: int \ 'a word)" + have "((=) ===> pcr_word) of_int of_int" by transfer_prover then show ?thesis by (simp add: id_def) qed +end + context semiring_1 begin lift_definition unsigned :: "'b::len0 word \ 'a" is "of_nat \ nat \ take_bit LENGTH('b)" by simp lemma unsigned_0 [simp]: "unsigned 0 = 0" by transfer simp end context semiring_char_0 begin lemma word_eq_iff_unsigned: "a = b \ unsigned a = unsigned b" by safe (transfer; simp add: eq_nat_nat_iff) end instantiation word :: (len0) equal begin definition equal_word :: "'a word \ 'a word \ bool" where "equal_word a b \ (unsigned a :: int) = unsigned b" instance proof fix a b :: "'a word" show "HOL.equal a b \ a = b" using word_eq_iff_unsigned [of a b] by (auto simp add: equal_word_def) qed end context ring_1 begin lift_definition signed :: "'b::len word \ 'a" is "of_int \ signed_take_bit (LENGTH('b) - 1)" by (simp add: signed_take_bit_eq_iff_take_bit_eq [symmetric]) lemma signed_0 [simp]: "signed 0 = 0" by transfer simp end lemma unsigned_of_nat [simp]: "unsigned (of_nat n :: 'a word) = take_bit LENGTH('a::len) n" by transfer (simp add: nat_eq_iff take_bit_eq_mod zmod_int) lemma of_nat_unsigned [simp]: "of_nat (unsigned a) = a" by transfer simp lemma of_int_unsigned [simp]: "of_int (unsigned a) = a" by transfer simp context ring_char_0 begin lemma word_eq_iff_signed: "a = b \ signed a = signed b" by safe (transfer; auto simp add: signed_take_bit_eq_iff_take_bit_eq) end lemma signed_of_int [simp]: "signed (of_int k :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) k" by transfer simp lemma of_int_signed [simp]: "of_int (signed a) = a" by transfer (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod mod_simps) subsubsection \Properties\ subsubsection \Division\ instantiation word :: (len0) modulo begin lift_definition divide_word :: "'a word \ 'a word \ 'a word" is "\a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b" by simp lift_definition modulo_word :: "'a word \ 'a word \ 'a word" is "\a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b" by simp instance .. end +context + includes lifting_syntax +begin + lemma [transfer_rule]: - "rel_fun pcr_word (\) even ((dvd) 2 :: 'a::len word \ bool)" + "(pcr_word ===> (\)) even ((dvd) 2 :: 'a::len word \ bool)" proof - have even_word_unfold: "even k \ (\l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \ ?Q") for k :: int proof assume ?P then show ?Q by auto next assume ?Q then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" .. then have "even (take_bit LENGTH('a) k)" by simp then show ?P by simp qed show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]) transfer_prover qed +end + instance word :: (len) semiring_modulo proof show "a div b * b + a mod b = a" for a b :: "'a word" proof transfer fix k l :: int define r :: int where "r = 2 ^ LENGTH('a)" then have r: "take_bit LENGTH('a) k = k mod r" for k by (simp add: take_bit_eq_mod) have "k mod r = ((k mod r) div (l mod r) * (l mod r) + (k mod r) mod (l mod r)) mod r" by (simp add: div_mult_mod_eq) also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_add_left_eq) also have "... = (((k mod r) div (l mod r) * l) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_mult_right_eq) finally have "k mod r = ((k mod r) div (l mod r) * l + (k mod r) mod (l mod r)) mod r" by (simp add: mod_simps) with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k" by simp qed qed instance word :: (len) semiring_parity proof show "\ 2 dvd (1::'a word)" by transfer simp consider (triv) "LENGTH('a) = 1" "take_bit LENGTH('a) 2 = (0 :: int)" | (take_bit_2) "take_bit LENGTH('a) 2 = (2 :: int)" proof (cases "LENGTH('a) \ 2") case False then have "LENGTH('a) = 1" by (auto simp add: not_le dest: less_2_cases) then have "take_bit LENGTH('a) 2 = (0 :: int)" by simp with \LENGTH('a) = 1\ triv show ?thesis by simp next case True then obtain n where "LENGTH('a) = Suc (Suc n)" by (auto dest: le_Suc_ex) then have "take_bit LENGTH('a) 2 = (2 :: int)" by simp with take_bit_2 show ?thesis by simp qed note * = this show even_iff_mod_2_eq_0: "2 dvd a \ a mod 2 = 0" for a :: "'a word" by (transfer; cases rule: *) (simp_all add: mod_2_eq_odd) show "\ 2 dvd a \ a mod 2 = 1" for a :: "'a word" by (transfer; cases rule: *) (simp_all add: mod_2_eq_odd) qed subsubsection \Orderings\ instantiation word :: (len0) linorder begin lift_definition less_eq_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a \ take_bit LENGTH('a) b" by simp lift_definition less_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b" by simp instance by standard (transfer; auto)+ end context linordered_semidom begin lemma word_less_eq_iff_unsigned: "a \ b \ unsigned a \ unsigned b" by (transfer fixing: less_eq) (simp add: nat_le_eq_zle) lemma word_less_iff_unsigned: "a < b \ unsigned a < unsigned b" by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative]) end subsection \Bit operation on \<^typ>\'a word\\ context unique_euclidean_semiring_with_nat begin primrec n_bits_of :: "nat \ 'a \ bool list" where "n_bits_of 0 a = []" | "n_bits_of (Suc n) a = odd a # n_bits_of n (a div 2)" lemma n_bits_of_eq_iff: "n_bits_of n a = n_bits_of n b \ take_bit n a = take_bit n b" apply (induction n arbitrary: a b) apply auto apply (metis local.dvd_add_times_triv_left_iff local.dvd_triv_right local.odd_one) apply (metis local.dvd_add_times_triv_left_iff local.dvd_triv_right local.odd_one) done lemma take_n_bits_of [simp]: "take m (n_bits_of n a) = n_bits_of (min m n) a" proof - define q and v and w where "q = min m n" and "v = m - q" and "w = n - q" then have "v = 0 \ w = 0" by auto then have "take (q + v) (n_bits_of (q + w) a) = n_bits_of q a" by (induction q arbitrary: a) auto with q_def v_def w_def show ?thesis by simp qed lemma unsigned_of_bits_n_bits_of [simp]: "unsigned_of_bits (n_bits_of n a) = take_bit n a" by (induction n arbitrary: a) (simp_all add: ac_simps) end lemma unsigned_of_bits_eq_of_bits: "unsigned_of_bits bs = (of_bits (bs @ [False]) :: int)" by (simp add: of_bits_int_def) instantiation word :: (len) bit_representation begin lift_definition bits_of_word :: "'a word \ bool list" is "n_bits_of LENGTH('a)" by (simp add: n_bits_of_eq_iff) lift_definition of_bits_word :: "bool list \ 'a word" is unsigned_of_bits . instance proof fix a :: "'a word" show "of_bits (bits_of a) = a" by transfer simp qed end lemma take_bit_complement_iff: "take_bit n (complement k) = take_bit n (complement l) \ take_bit n k = take_bit n l" for k l :: int by (simp add: take_bit_eq_mod mod_eq_dvd_iff dvd_diff_commute) lemma take_bit_not_iff: "take_bit n (NOT k) = take_bit n (NOT l) \ take_bit n k = take_bit n l" for k l :: int by (simp add: not_int_def take_bit_complement_iff) lemma n_bits_of_not: "n_bits_of n (NOT k) = map Not (n_bits_of n k)" for k :: int by (induction n arbitrary: k) (simp_all add: not_div_2) lemma take_bit_and [simp]: "take_bit n (k AND l) = take_bit n k AND take_bit n l" for k l :: int apply (induction n arbitrary: k l) apply simp apply (subst and_int.rec) apply (subst (2) and_int.rec) apply simp done lemma take_bit_or [simp]: "take_bit n (k OR l) = take_bit n k OR take_bit n l" for k l :: int apply (induction n arbitrary: k l) apply simp apply (subst or_int.rec) apply (subst (2) or_int.rec) apply simp done lemma take_bit_xor [simp]: "take_bit n (k XOR l) = take_bit n k XOR take_bit n l" for k l :: int apply (induction n arbitrary: k l) apply simp apply (subst xor_int.rec) apply (subst (2) xor_int.rec) apply simp done instantiation word :: (len) bit_operations begin lift_definition not_word :: "'a word \ 'a word" is not by (simp add: take_bit_not_iff) lift_definition and_word :: "'a word \ 'a word \ 'a word" is "and" by simp lift_definition or_word :: "'a word \ 'a word \ 'a word" is or by simp lift_definition xor_word :: "'a word \ 'a word \ 'a word" is xor by simp lift_definition shift_left_word :: "'a word \ nat \ 'a word" is shift_left proof - show "take_bit LENGTH('a) (k << n) = take_bit LENGTH('a) (l << n)" if "take_bit LENGTH('a) k = take_bit LENGTH('a) l" for k l :: int and n :: nat proof - from that have "take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k) = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)" by simp moreover have "min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n" by simp ultimately show ?thesis by (simp add: take_bit_push_bit) qed qed lift_definition shift_right_word :: "'a word \ nat \ 'a word" is "\k n. drop_bit n (take_bit LENGTH('a) k)" by simp instance proof show "semilattice ((AND) :: 'a word \ _)" by standard (transfer; simp add: ac_simps)+ show "semilattice ((OR) :: 'a word \ _)" by standard (transfer; simp add: ac_simps)+ show "abel_semigroup ((XOR) :: 'a word \ _)" by standard (transfer; simp add: ac_simps)+ show "not = (of_bits \ map Not \ bits_of :: 'a word \ 'a word)" proof fix a :: "'a word" have "NOT a = of_bits (map Not (bits_of a))" by transfer (simp flip: unsigned_of_bits_take n_bits_of_not add: take_map) then show "NOT a = (of_bits \ map Not \ bits_of) a" by simp qed show "of_bits bs AND of_bits cs = (of_bits (map2 (\) bs cs) :: 'a word)" if "length bs = length cs" for bs cs using that apply transfer apply (simp only: unsigned_of_bits_eq_of_bits) apply (subst and_eq) apply simp_all done show "of_bits bs OR of_bits cs = (of_bits (map2 (\) bs cs) :: 'a word)" if "length bs = length cs" for bs cs using that apply transfer apply (simp only: unsigned_of_bits_eq_of_bits) apply (subst or_eq) apply simp_all done show "of_bits bs XOR of_bits cs = (of_bits (map2 (\) bs cs) :: 'a word)" if "length bs = length cs" for bs cs using that apply transfer apply (simp only: unsigned_of_bits_eq_of_bits) apply (subst xor_eq) apply simp_all done show "a << n = of_bits (replicate n False @ bits_of a)" for a :: "'a word" and n :: nat by transfer (simp add: push_bit_take_bit) show "a >> n = of_bits (drop n (bits_of a))" if "n < length (bits_of a)" for a :: "'a word" and n :: nat using that by transfer simp qed end global_interpretation bit_word: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: 'a::len word" rewrites "bit_word.xor = ((XOR) :: 'a word \ _)" proof - interpret bit_word: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: 'a word" proof show "a AND (b OR c) = a AND b OR a AND c" for a b c :: "'a word" by transfer (simp add: bit_int.conj_disj_distrib) show "a OR b AND c = (a OR b) AND (a OR c)" for a b c :: "'a word" by transfer (simp add: bit_int.disj_conj_distrib) show "a AND NOT a = 0" for a :: "'a word" by transfer simp show "a OR NOT a = - 1" for a :: "'a word" by transfer simp qed (transfer; simp)+ show "boolean_algebra (AND) (OR) NOT 0 (- 1 :: 'a word)" by (fact bit_word.boolean_algebra_axioms) show "bit_word.xor = ((XOR) :: 'a word \ _)" proof (rule ext)+ fix a b :: "'a word" have "a XOR b = a AND NOT b OR NOT a AND b" by transfer (simp add: bit_int.xor_def) then show "bit_word.xor a b = a XOR b" by (simp add: bit_word.xor_def) qed qed end