# HG changeset patch # User Christian Urban # Date 1260138902 -3600 # Node ID 97f6e5c61f030132dd31f94ba1827ed48ebbd798 # Parent 7414f6cb539868dbea1d4ef8e87f6bed03d4db9e# Parent 2da4fb1894d2d012301475182fce2a9da32ef970 merged diff -r 7414f6cb5398 -r 97f6e5c61f03 IntEx2.thy --- a/IntEx2.thy Sun Dec 06 23:32:27 2009 +0100 +++ b/IntEx2.thy Sun Dec 06 23:35:02 2009 +0100 @@ -54,7 +54,7 @@ where "minus_raw (x, y) = (y, x)" -quotient_def +quotient_def minus_qnt::"int \ int" where "minus_qnt \ minus_raw" @@ -111,6 +111,10 @@ shows "(op \ ===> op \ ===> op \) plus_raw plus_raw" by auto +lemma minus_raw_rsp[quotient_rsp]: + shows "(op \ ===> op \) minus_raw minus_raw" + by auto + lemma mult_raw_rsp[quotient_rsp]: shows "(op \ ===> op \ ===> op \) mult_raw mult_raw" apply(auto) @@ -156,9 +160,9 @@ (simp add: mult algebra_simps) lemma one_zero_distinct: - shows "(0, 0) \ ((1::nat), (0::nat))" + shows "\ (0, 0) \ ((1::nat), (0::nat))" by simp - + text{*The integers form a @{text comm_ring_1}*} @@ -171,69 +175,42 @@ 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 *}) + apply(tactic {* lift_tac @{context} @{thm plus_assoc_raw} [@{thm int_equivp}] 1 *}) done show "i + j = j + i" unfolding add_int_def - apply(tactic {* lift_tac @{context} @{thm plus_sym_raw} 1 *}) + apply(tactic {* lift_tac @{context} @{thm plus_sym_raw} [@{thm int_equivp}] 1 *}) done show "0 + i = (i::int)" unfolding add_int_def Zero_int_def - apply(tactic {* procedure_tac @{context} @{thm plus_zero_raw} 1 *}) - apply(tactic {* regularize_tac @{context} 1 *}) - apply(tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) - defer - apply(tactic {* clean_tac @{context} 1*}) - sorry + apply(tactic {* lift_tac @{context} @{thm plus_zero_raw} [@{thm int_equivp}] 1 *}) + done show "- i + i = 0" unfolding add_int_def minus_int_def Zero_int_def - apply(tactic {* procedure_tac @{context} @{thm plus_minus_zero_raw} 1 *}) - apply(tactic {* regularize_tac @{context} 1 *}) - apply(tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) - defer - apply(tactic {* clean_tac @{context} 1*}) - sorry + apply(tactic {* lift_tac @{context} @{thm plus_minus_zero_raw} [@{thm int_equivp}] 1 *}) + done show "i - j = i + - j" by (simp add: diff_int_def) show "(i * j) * k = i * (j * k)" unfolding mult_int_def - apply(tactic {* procedure_tac @{context} @{thm mult_assoc_raw} 1 *}) - apply(tactic {* regularize_tac @{context} 1 *}) - apply(tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) - defer - apply(tactic {* clean_tac @{context} 1*}) - sorry + apply(tactic {* lift_tac @{context} @{thm mult_assoc_raw} [@{thm int_equivp}] 1 *}) + done show "i * j = j * i" unfolding mult_int_def - apply(tactic {* procedure_tac @{context} @{thm mult_sym_raw} 1 *}) - apply(tactic {* regularize_tac @{context} 1 *}) - apply(tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) - defer - apply(tactic {* clean_tac @{context} 1*}) - sorry + apply(tactic {* lift_tac @{context} @{thm mult_sym_raw} [@{thm int_equivp}] 1 *}) + done show "1 * i = i" unfolding mult_int_def One_int_def - apply(tactic {* procedure_tac @{context} @{thm mult_one_raw} 1 *}) - apply(tactic {* regularize_tac @{context} 1 *}) - apply(tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) - defer - apply(tactic {* clean_tac @{context} 1*}) - sorry + apply(tactic {* lift_tac @{context} @{thm mult_one_raw} [@{thm int_equivp}] 1 *}) + done show "(i + j) * k = i * k + j * k" unfolding mult_int_def add_int_def - apply(tactic {* procedure_tac @{context} @{thm mult_plus_comm_raw} 1 *}) - apply(tactic {* regularize_tac @{context} 1 *}) - apply(tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) - defer - apply(tactic {* clean_tac @{context} 1*}) - sorry + apply(tactic {* lift_tac @{context} @{thm mult_plus_comm_raw} [@{thm int_equivp}] 1 *}) + done show "0 \ (1::int)" unfolding Zero_int_def One_int_def - apply(tactic {* procedure_tac @{context} @{thm one_zero_distinct} 1 *}) - defer - defer - apply(tactic {* clean_tac @{context} 1*}) - sorry + apply(tactic {* lift_tac @{context} @{thm one_zero_distinct} [@{thm int_equivp}] 1 *}) + done qed term of_nat @@ -269,21 +246,21 @@ 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 *}) + apply(tactic {* lift_tac @{context} @{thm le_antisym_raw} [@{thm int_equivp}] 1 *}) done 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 *}) + apply(tactic {* lift_tac @{context} @{thm le_refl_raw} [@{thm int_equivp}] 1 *}) done show "i \ j \ j \ k \ i \ k" unfolding le_int_def - apply(tactic {* lift_tac @{context} @{thm le_trans_raw} 1 *}) + apply(tactic {* lift_tac @{context} @{thm le_trans_raw} [@{thm int_equivp}] 1 *}) done show "i \ j \ j \ i" unfolding le_int_def - apply(tactic {* lift_tac @{context} @{thm le_cases_raw} 1 *}) + apply(tactic {* lift_tac @{context} @{thm le_cases_raw} [@{thm int_equivp}] 1 *}) done qed @@ -312,13 +289,13 @@ 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 *}) + apply(tactic {* lift_tac @{context} @{thm le_plus_raw} [@{thm int_equivp}] 1 *}) done qed lemma test: - "\le_raw i j \ i \ j; le_raw (0, 0) k \ (0, 0) \ k\ - \ le_raw (mult_raw k i) (mult_raw k j) \ mult_raw k i \ mult_raw k j" + "\le_raw i j \ \i \ j; le_raw (0, 0) k \ \(0, 0) \ k\ + \ le_raw (mult_raw k i) (mult_raw k j) \ \mult_raw k i \ mult_raw k j" apply(cases i, cases j, cases k) apply(auto simp add: mult algebra_simps) sorry @@ -330,12 +307,8 @@ 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 {* procedure_tac @{context} @{thm test} 1 *}) - defer - apply(tactic {* inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) - defer - apply(tactic {* clean_tac @{context} 1*}) - sorry + apply(tactic {* lift_tac @{context} @{thm test} [@{thm int_equivp}] 1 *}) + done 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 0P1 TYP; - \ty name kind. \P2 ty; P1 kind\ \ P1 (KPI ty name kind); - \id. P2 (TCONST id); - \ty trm. \P2 ty; P3 trm\ \ P2 (TAPP ty trm); - \ty1 name ty2. \P2 ty1; P2 ty2\ \ P2 (TPI ty1 name ty2); - \id. P3 (CONS id); \name. P3 (VAR name); - \trm1 trm2. \P3 trm1; P3 trm2\ \ P3 (APP trm1 trm2); - \ty name trm. \P2 ty; P3 trm\ \ P3 (LAM ty name trm)\ - \ P1 mkind \ P2 mty \ P3 mtrm" -apply(tactic {* lift_tac @{context} @{thm kind_ty_trm.induct} 1 *}) +lemma "\P TYP; + \ty name kind. \Q ty; P kind\ \ P (KPI ty name kind); + \id. Q (TCONST id); + \ty trm. \Q ty; R trm\ \ Q (TAPP ty trm); + \ty1 name ty2. \Q ty1; Q ty2\ \ Q (TPI ty1 name ty2); + \id. R (CONS id); \name. R (VAR name); + \trm1 trm2. \R trm1; R trm2\ \ R (APP trm1 trm2); + \ty name trm. \Q ty; R trm\ \ R (LAM ty name trm)\ + \ P mkind \ Q mty \ R mtrm" +apply(tactic {* lift_tac @{context} @{thm kind_ty_trm.induct} @{thms alpha_equivps} 1 *}) done print_quotients