diff -r fdccdc52c68a -r 02fd9de9d45e Quot/Examples/IntEx2.thy --- a/Quot/Examples/IntEx2.thy Tue Dec 08 23:32:54 2009 +0100 +++ b/Quot/Examples/IntEx2.thy Wed Dec 09 00:03:18 2009 +0100 @@ -105,8 +105,6 @@ end -thm add_assoc - lemma plus_raw_rsp[quot_respect]: shows "(op \ ===> op \ ===> op \) plus_raw plus_raw" by auto @@ -171,42 +169,33 @@ fix i j k :: int show "(i + j) + k = i + (j + k)" unfolding add_int_def - apply(tactic {* lift_tac @{context} @{thm plus_assoc_raw} 1 *}) - done + by (lifting plus_assoc_raw) show "i + j = j + i" unfolding add_int_def - apply(tactic {* lift_tac @{context} @{thm plus_sym_raw} 1 *}) - done + by (lifting plus_sym_raw) show "0 + i = (i::int)" unfolding add_int_def Zero_int_def - apply(tactic {* lift_tac @{context} @{thm plus_zero_raw} 1 *}) - done + by (lifting plus_zero_raw) show "- i + i = 0" unfolding add_int_def minus_int_def Zero_int_def - apply(tactic {* lift_tac @{context} @{thm plus_minus_zero_raw} 1 *}) - done + by (lifting plus_minus_zero_raw) show "i - j = i + - j" by (simp add: diff_int_def) show "(i * j) * k = i * (j * k)" unfolding mult_int_def - apply(tactic {* lift_tac @{context} @{thm mult_assoc_raw} 1 *}) - done + by (lifting mult_assoc_raw) show "i * j = j * i" unfolding mult_int_def - apply(tactic {* lift_tac @{context} @{thm mult_sym_raw} 1 *}) - done + by (lifting mult_sym_raw) show "1 * i = i" unfolding mult_int_def One_int_def - apply(tactic {* lift_tac @{context} @{thm mult_one_raw} 1 *}) - done + by (lifting mult_one_raw) show "(i + j) * k = i * k + j * k" unfolding mult_int_def add_int_def - apply(tactic {* lift_tac @{context} @{thm mult_plus_comm_raw} 1 *}) - done + by (lifting mult_plus_comm_raw) show "0 \ (1::int)" unfolding Zero_int_def One_int_def - apply(tactic {* lift_tac @{context} @{thm one_zero_distinct} 1 *}) - done + by (lifting one_zero_distinct) qed term of_nat @@ -242,22 +231,18 @@ fix i j k :: int show antisym: "i \ j \ j \ i \ i = j" unfolding le_int_def - apply(tactic {* lift_tac @{context} @{thm le_antisym_raw} 1 *}) - done + by (lifting le_antisym_raw) show "(i < j) = (i \ j \ \ j \ i)" by (auto simp add: less_int_def dest: antisym) show "i \ i" unfolding le_int_def - apply(tactic {* lift_tac @{context} @{thm le_refl_raw} 1 *}) - done + by (lifting le_refl_raw) show "i \ j \ j \ k \ i \ k" unfolding le_int_def - apply(tactic {* lift_tac @{context} @{thm le_trans_raw} 1 *}) - done + by (lifting le_trans_raw) show "i \ j \ j \ i" unfolding le_int_def - apply(tactic {* lift_tac @{context} @{thm le_cases_raw} 1 *}) - done + by (lifting le_cases_raw) qed instantiation int :: distrib_lattice @@ -285,8 +270,7 @@ fix i j k :: int show "i \ j \ k + i \ k + j" unfolding le_int_def add_int_def - apply(tactic {* lift_tac @{context} @{thm le_plus_raw} 1 *}) - done + by (lifting le_plus_raw) qed lemma test: @@ -303,8 +287,7 @@ fix i j k :: int show "i < j \ 0 < k \ k * i < k * j" unfolding mult_int_def le_int_def less_int_def Zero_int_def - apply(tactic {* lift_tac @{context} @{thm test} 1 *}) - done + by (lifting test) 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