diff -r 1c0dfa7ebb72 -r 3f0ab2c306f7 src/HOL/ex/Mutil.ML --- a/src/HOL/ex/Mutil.ML Wed May 07 13:50:52 1997 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,172 +0,0 @@ -(* Title: HOL/ex/Mutil - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1996 University of Cambridge - -The Mutilated Chess Board Problem, formalized inductively -*) - -open Mutil; - -Addsimps tiling.intrs; - -(** The union of two disjoint tilings is a tiling **) - -goal thy "!!t. t: tiling A ==> \ -\ u: tiling A --> t <= Compl u --> t Un u : tiling A"; -by (etac tiling.induct 1); -by (Simp_tac 1); -by (simp_tac (!simpset addsimps [Un_assoc]) 1); -by (blast_tac (!claset addIs tiling.intrs) 1); -qed_spec_mp "tiling_UnI"; - - -(*** Chess boards ***) - -val [below_0, below_Suc] = nat_recs below_def; -Addsimps [below_0, below_Suc]; - -goal thy "ALL i. (i: below k) = (i finite(evnodd X b) *) -bind_thm("finite_evnodd", evnodd_subset RS finite_subset); - -goalw thy [evnodd_def] "evnodd (A Un B) b = evnodd A b Un evnodd B b"; -by (Blast_tac 1); -qed "evnodd_Un"; - -goalw thy [evnodd_def] "evnodd (A - B) b = evnodd A b - evnodd B b"; -by (Blast_tac 1); -qed "evnodd_Diff"; - -goalw thy [evnodd_def] "evnodd {} b = {}"; -by (Simp_tac 1); -qed "evnodd_empty"; - -goalw thy [evnodd_def] - "evnodd (insert (i,j) C) b = \ -\ (if (i+j) mod 2 = b then insert (i,j) (evnodd C b) else evnodd C b)"; -by (asm_full_simp_tac (!simpset addsimps [evnodd_def] - setloop (split_tac [expand_if] THEN' Step_tac)) 1); -qed "evnodd_insert"; - - -(*** Dominoes ***) - -goal thy "!!d. [| d:domino; b<2 |] ==> EX i j. evnodd d b = {(i,j)}"; -by (eresolve_tac [domino.elim] 1); -by (res_inst_tac [("k1", "i+j")] (mod2_cases RS disjE) 2); -by (res_inst_tac [("k1", "i+j")] (mod2_cases RS disjE) 1); -by (REPEAT_FIRST assume_tac); -(*Four similar cases: case (i+j) mod 2 = b, 2#-b, ...*) -by (REPEAT (asm_full_simp_tac (!simpset addsimps - [less_Suc_eq, evnodd_insert, evnodd_empty, mod_Suc] - setloop split_tac [expand_if]) 1 - THEN Blast_tac 1)); -qed "domino_singleton"; - -goal thy "!!d. d:domino ==> finite d"; -by (blast_tac (!claset addSIs [finite_insertI, finite_emptyI] - addSEs [domino.elim]) 1); -qed "domino_finite"; - - -(*** Tilings of dominoes ***) - -goal thy "!!t. t:tiling domino ==> finite t"; -by (eresolve_tac [tiling.induct] 1); -by (rtac finite_emptyI 1); -by (blast_tac (!claset addSIs [finite_UnI] addIs [domino_finite]) 1); -qed "tiling_domino_finite"; - -goal thy "!!t. t: tiling domino ==> card(evnodd t 0) = card(evnodd t 1)"; -by (eresolve_tac [tiling.induct] 1); -by (simp_tac (!simpset addsimps [evnodd_def]) 1); -by (res_inst_tac [("b1","0")] (domino_singleton RS exE) 1); -by (Simp_tac 2 THEN assume_tac 1); -by (res_inst_tac [("b1","1")] (domino_singleton RS exE) 1); -by (Simp_tac 2 THEN assume_tac 1); -by (Step_tac 1); -by (subgoal_tac "ALL p b. p : evnodd a b --> p ~: evnodd ta b" 1); -by (asm_simp_tac (!simpset addsimps [evnodd_Un, Un_insert_left, - tiling_domino_finite, - evnodd_subset RS finite_subset, - card_insert_disjoint]) 1); -by (blast_tac (!claset addSDs [evnodd_subset RS subsetD] addEs [equalityE]) 1); -qed "tiling_domino_0_1"; - -goal thy "!!m n. [| t = below(Suc m + Suc m) Times below(Suc n + Suc n); \ -\ t' = t - {(0,0)} - {(Suc(m+m), Suc(n+n))} \ -\ |] ==> t' ~: tiling domino"; -by (rtac notI 1); -by (dtac tiling_domino_0_1 1); -by (subgoal_tac "card(evnodd t' 0) < card(evnodd t' 1)" 1); -by (Asm_full_simp_tac 1); -by (subgoal_tac "t : tiling domino" 1); -(*Requires a small simpset that won't move the Suc applications*) -by (asm_simp_tac (HOL_ss addsimps [dominoes_tile_matrix]) 2); -by (subgoal_tac "(m+m)+(n+n) = (m+n)+(m+n)" 1); -by (asm_simp_tac (!simpset addsimps add_ac) 2); -by (asm_full_simp_tac - (!simpset addsimps [evnodd_Diff, evnodd_insert, evnodd_empty, - mod_less, tiling_domino_0_1 RS sym]) 1); -by (rtac less_trans 1); -by (REPEAT - (rtac card_Diff 1 - THEN - asm_simp_tac (!simpset addsimps [tiling_domino_finite, finite_evnodd]) 1 - THEN - asm_simp_tac (!simpset addsimps [mod_less, evnodd_iff]) 1)); -qed "mutil_not_tiling"; -