diff -r db158e995bfc -r 9df6144e281b Attic/Quot/Examples/IntEx2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Attic/Quot/Examples/IntEx2.thy Thu Feb 25 07:57:17 2010 +0100 @@ -0,0 +1,445 @@ +theory IntEx2 +imports "../Quotient" "../Quotient_Product" Nat +(*uses + ("Tools/numeral.ML") + ("Tools/numeral_syntax.ML") + ("Tools/int_arith.ML")*) +begin + +fun + intrel :: "(nat \ nat) \ (nat \ nat) \ bool" (infix "\" 50) +where + "intrel (x, y) (u, v) = (x + v = u + y)" + +quotient_type int = "nat \ nat" / intrel + unfolding equivp_def + by (auto simp add: mem_def expand_fun_eq) + +instantiation int :: "{zero, one, plus, uminus, minus, times, ord, abs, sgn}" +begin + +ML {* @{term "0 \ int"} *} + +quotient_definition + "0 \ int" is "(0\nat, 0\nat)" + +quotient_definition + "1 \ int" is "(1\nat, 0\nat)" + +fun + plus_raw :: "(nat \ nat) \ (nat \ nat) \ (nat \ nat)" +where + "plus_raw (x, y) (u, v) = (x + u, y + v)" + +quotient_definition + "(op +) \ (int \ int \ int)" is "plus_raw" + +fun + uminus_raw :: "(nat \ nat) \ (nat \ nat)" +where + "uminus_raw (x, y) = (y, x)" + +quotient_definition + "(uminus \ (int \ int))" is "uminus_raw" + +definition + minus_int_def [code del]: "z - w = z + (-w\int)" + +fun + mult_raw :: "(nat \ nat) \ (nat \ nat) \ (nat \ nat)" +where + "mult_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)" + +quotient_definition + mult_int_def: "(op *) :: (int \ int \ int)" is "mult_raw" + +fun + le_raw :: "(nat \ nat) \ (nat \ nat) \ bool" +where + "le_raw (x, y) (u, v) = (x+v \ u+y)" + +quotient_definition + le_int_def: "(op \) :: int \ int \ bool" is "le_raw" + +definition + less_int_def [code del]: "(z\int) < w = (z \ w \ z \ w)" + +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 .. + +end + +lemma plus_raw_rsp[quot_respect]: + shows "(op \ ===> op \ ===> op \) plus_raw plus_raw" +by auto + +lemma uminus_raw_rsp[quot_respect]: + shows "(op \ ===> op \) uminus_raw uminus_raw" + by auto + +lemma mult_raw_fst: + assumes a: "x \ z" + shows "mult_raw x y \ mult_raw z y" +using a +apply(cases x, cases y, cases z) +apply(auto simp add: mult_raw.simps intrel.simps) +apply(rename_tac u v w x y z) +apply(subgoal_tac "u*w + z*w = y*w + v*w & u*x + z*x = y*x + v*x") +apply(simp add: mult_ac) +apply(simp add: add_mult_distrib [symmetric]) +done + +lemma mult_raw_snd: + assumes a: "x \ z" + shows "mult_raw y x \ mult_raw y z" +using a +apply(cases x, cases y, cases z) +apply(auto simp add: mult_raw.simps intrel.simps) +apply(rename_tac u v w x y z) +apply(subgoal_tac "u*w + z*w = y*w + v*w & u*x + z*x = y*x + v*x") +apply(simp add: mult_ac) +apply(simp add: add_mult_distrib [symmetric]) +done + +lemma mult_raw_rsp[quot_respect]: + shows "(op \ ===> op \ ===> op \) mult_raw mult_raw" +apply(simp only: fun_rel_def) +apply(rule allI | rule impI)+ +apply(rule equivp_transp[OF int_equivp]) +apply(rule mult_raw_fst) +apply(assumption) +apply(rule mult_raw_snd) +apply(assumption) +done + +lemma le_raw_rsp[quot_respect]: + shows "(op \ ===> op \ ===> op =) le_raw le_raw" +by auto + +lemma plus_assoc_raw: + shows "plus_raw (plus_raw i j) k \ plus_raw i (plus_raw j k)" +by (cases i, cases j, cases k) (simp) + +lemma plus_sym_raw: + shows "plus_raw i j \ plus_raw j i" +by (cases i, cases j) (simp) + +lemma plus_zero_raw: + shows "plus_raw (0, 0) i \ i" +by (cases i) (simp) + +lemma plus_minus_zero_raw: + shows "plus_raw (uminus_raw i) i \ (0, 0)" +by (cases i) (simp) + +lemma times_assoc_raw: + shows "mult_raw (mult_raw i j) k \ mult_raw i (mult_raw j k)" +by (cases i, cases j, cases k) + (simp add: algebra_simps) + +lemma times_sym_raw: + shows "mult_raw i j \ mult_raw j i" +by (cases i, cases j) (simp add: algebra_simps) + +lemma times_one_raw: + shows "mult_raw (1, 0) i \ i" +by (cases i) (simp) + +lemma times_plus_comm_raw: + shows "mult_raw (plus_raw i j) k \ plus_raw (mult_raw i k) (mult_raw j k)" +by (cases i, cases j, cases k) + (simp add: algebra_simps) + +lemma one_zero_distinct: + shows "\ (0, 0) \ ((1::nat), (0::nat))" + by simp + +text{* The integers form a @{text comm_ring_1}*} + +instance int :: comm_ring_1 +proof + fix i j k :: int + show "(i + j) + k = i + (j + k)" + by (lifting plus_assoc_raw) + show "i + j = j + i" + by (lifting plus_sym_raw) + show "0 + i = (i::int)" + by (lifting plus_zero_raw) + show "- i + i = 0" + by (lifting plus_minus_zero_raw) + show "i - j = i + - j" + by (simp add: minus_int_def) + show "(i * j) * k = i * (j * k)" + by (lifting times_assoc_raw) + show "i * j = j * i" + by (lifting times_sym_raw) + show "1 * i = i" + by (lifting times_one_raw) + show "(i + j) * k = i * k + j * k" + by (lifting times_plus_comm_raw) + show "0 \ (1::int)" + by (lifting one_zero_distinct) +qed + +lemma plus_raw_rsp_aux: + assumes a: "a \ b" "c \ d" + shows "plus_raw a c \ plus_raw b d" +using a +by (cases a, cases b, cases c, cases d) + (simp) + +lemma add: + "(abs_int (x,y)) + (abs_int (u,v)) = + (abs_int (x + u, y + v))" +apply(simp add: plus_int_def id_simps) +apply(fold plus_raw.simps) +apply(rule Quotient_rel_abs[OF Quotient_int]) +apply(rule plus_raw_rsp_aux) +apply(simp_all add: rep_abs_rsp_left[OF Quotient_int]) +done + +definition int_of_nat_raw: + "int_of_nat_raw m = (m :: nat, 0 :: nat)" + +quotient_definition + "int_of_nat :: nat \ int" is "int_of_nat_raw" + +lemma[quot_respect]: + shows "(op = ===> op \) int_of_nat_raw int_of_nat_raw" +by (simp add: equivp_reflp[OF int_equivp]) + +lemma int_of_nat: + shows "of_nat m = int_of_nat m" +apply (induct m) +apply (simp_all add: zero_int_def one_int_def int_of_nat_def int_of_nat_raw add) +done + +lemma le_antisym_raw: + shows "le_raw i j \ le_raw j i \ i \ j" +by (cases i, cases j) (simp) + +lemma le_refl_raw: + shows "le_raw i i" +by (cases i) (simp) + +lemma le_trans_raw: + shows "le_raw i j \ le_raw j k \ le_raw i k" +by (cases i, cases j, cases k) (simp) + +lemma le_cases_raw: + shows "le_raw i j \ le_raw j i" +by (cases i, cases j) + (simp add: linorder_linear) + +instance int :: linorder +proof + fix i j k :: int + show antisym: "i \ j \ j \ i \ i = j" + by (lifting le_antisym_raw) + show "(i < j) = (i \ j \ \ j \ i)" + by (auto simp add: less_int_def dest: antisym) + show "i \ i" + by (lifting le_refl_raw) + show "i \ j \ j \ k \ i \ k" + by (lifting le_trans_raw) + show "i \ j \ j \ i" + by (lifting le_cases_raw) +qed + +instantiation int :: distrib_lattice +begin + +definition + "(inf \ int \ int \ int) = min" + +definition + "(sup \ int \ int \ int) = max" + +instance + by intro_classes + (auto simp add: inf_int_def sup_int_def min_max.sup_inf_distrib1) + +end + +lemma le_plus_raw: + shows "le_raw i j \ le_raw (plus_raw k i) (plus_raw k j)" +by (cases i, cases j, cases k) (simp) + + +instance int :: ordered_cancel_ab_semigroup_add +proof + fix i j k :: int + show "i \ j \ k + i \ k + j" + by (lifting le_plus_raw) +qed + +abbreviation + "less_raw i j \ le_raw i j \ \(i \ j)" + +lemma zmult_zless_mono2_lemma: + fixes i j::int + and k::nat + shows "i < j \ 0 < k \ of_nat k * i < of_nat k * j" +apply(induct "k") +apply(simp) +apply(case_tac "k = 0") +apply(simp_all add: left_distrib add_strict_mono) +done + +lemma zero_le_imp_eq_int_raw: + fixes k::"(nat \ nat)" + shows "less_raw (0, 0) k \ (\n > 0. k \ int_of_nat_raw n)" +apply(cases k) +apply(simp add:int_of_nat_raw) +apply(auto) +apply(rule_tac i="b" and j="a" in less_Suc_induct) +apply(auto) +done + +lemma zero_le_imp_eq_int: + fixes k::int + shows "0 < k \ \n > 0. k = of_nat n" + unfolding less_int_def int_of_nat + by (lifting zero_le_imp_eq_int_raw) + +lemma zmult_zless_mono2: + fixes i j k::int + assumes a: "i < j" "0 < k" + shows "k * i < k * j" +using a +using a +apply(drule_tac zero_le_imp_eq_int) +apply(auto simp add: zmult_zless_mono2_lemma) +done + +text{*The integers form an ordered integral domain*} +instance int :: linordered_idom +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 'a" +where + "of_int +*) + + +subsection {* Binary representation *} + +text {* + This formalization defines binary arithmetic in terms of the integers + rather than using a datatype. This avoids multiple representations (leading + zeroes, etc.) See @{text "ZF/Tools/twos-compl.ML"}, function @{text + int_of_binary}, for the numerical interpretation. + + The representation expects that @{text "(m mod 2)"} is 0 or 1, + even if m is negative; + For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus + @{text "-5 = (-3)*2 + 1"}. + + This two's complement binary representation derives from the paper + "An Efficient Representation of Arithmetic for Term Rewriting" by + Dave Cohen and Phil Watson, Rewriting Techniques and Applications, + Springer LNCS 488 (240-251), 1991. +*} + +subsubsection {* The constructors @{term Bit0}, @{term Bit1}, @{term Pls} and @{term Min} *} + +definition + Pls :: int where + [code del]: "Pls = 0" + +definition + Min :: int where + [code del]: "Min = - 1" + +definition + Bit0 :: "int \ int" where + [code del]: "Bit0 k = k + k" + +definition + Bit1 :: "int \ int" where + [code del]: "Bit1 k = 1 + k + k" + +class number = -- {* for numeric types: nat, int, real, \dots *} + fixes number_of :: "int \ 'a" + +(*use "~~/src/HOL/Tools/numeral.ML" + +syntax + "_Numeral" :: "num_const \ 'a" ("_") + +use "~~/src/HOL/Tools/numeral_syntax.ML" + +setup NumeralSyntax.setup + +abbreviation + "Numeral0 \ number_of Pls" + +abbreviation + "Numeral1 \ number_of (Bit1 Pls)" + +lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)" + -- {* Unfold all @{text let}s involving constants *} + unfolding Let_def .. + +definition + succ :: "int \ int" where + [code del]: "succ k = k + 1" + +definition + pred :: "int \ int" where + [code del]: "pred k = k - 1" + +lemmas + max_number_of [simp] = max_def + [of "number_of u" "number_of v", standard, simp] +and + min_number_of [simp] = min_def + [of "number_of u" "number_of v", standard, simp] + -- {* unfolding @{text minx} and @{text max} on numerals *} + +lemmas numeral_simps = + succ_def pred_def Pls_def Min_def Bit0_def Bit1_def + +text {* Removal of leading zeroes *} + +lemma Bit0_Pls [simp, code_post]: + "Bit0 Pls = Pls" + unfolding numeral_simps by simp + +lemma Bit1_Min [simp, code_post]: + "Bit1 Min = Min" + unfolding numeral_simps by simp + +lemmas normalize_bin_simps = + Bit0_Pls Bit1_Min +*) + +end