merged
authorhuffman
Thu, 04 Mar 2010 18:18:52 -0800
changeset 35575 374c638a796e
parent 35574 ee5df989b7c4 (diff)
parent 35571 4af56a1c4c7d (current diff)
child 35583 c7ddb7105dde
merged
--- a/src/HOLCF/FOCUS/Fstreams.thy	Thu Mar 04 22:48:50 2010 +0100
+++ b/src/HOLCF/FOCUS/Fstreams.thy	Thu Mar 04 18:18:52 2010 -0800
@@ -240,7 +240,6 @@
      ==> (EX j t. Y j = <a> ooo t) & (EX X. chain X & (ALL i. EX j. <a> ooo X i << Y j) & (LUB i. X i) = s)"
 apply (auto simp add: fstreams_lub_lemma1)
 apply (rule_tac x="%n. stream_take n$s" in exI, auto)
-apply (simp add: chain_stream_take)
 apply (induct_tac i, auto)
 apply (drule fstreams_lub_lemma1, auto)
 apply (rule_tac x="j" in exI, auto)
@@ -293,7 +292,6 @@
       ==> (EX j t. Y j = (a, <m> ooo t)) & (EX X. chain X & (ALL i. EX j. (a, <m> ooo X i) << Y j) & (LUB i. X i) = ms)"
 apply (auto simp add: fstreams_lub_lemma2)
 apply (rule_tac x="%n. stream_take n$ms" in exI, auto)
-apply (simp add: chain_stream_take)
 apply (induct_tac i, auto)
 apply (drule fstreams_lub_lemma2, auto)
 apply (rule_tac x="j" in exI, auto)
--- a/src/HOLCF/Tools/Domain/domain_take_proofs.ML	Thu Mar 04 22:48:50 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_take_proofs.ML	Thu Mar 04 18:18:52 2010 -0800
@@ -188,6 +188,13 @@
         (Thm.no_attributes (Binding.name name, thm))
     ||> Sign.parent_path;
 
+fun add_qualified_simp_thm name (path, thm) thy =
+    thy
+    |> Sign.add_path path
+    |> yield_singleton PureThy.add_thms
+        ((Binding.name name, thm), [Simplifier.simp_add])
+    ||> Sign.parent_path;
+
 (******************************************************************************)
 (************************** defining take functions ***************************)
 (******************************************************************************)
@@ -262,9 +269,9 @@
         val goal = mk_trp (mk_chain take_const);
         val rules = take_defs @ @{thms chain_iterate ch2ch_fst ch2ch_snd};
         val tac = simp_tac (HOL_basic_ss addsimps rules) 1;
-        val chain_take_thm = Goal.prove_global thy [] [] goal (K tac);
+        val thm = Goal.prove_global thy [] [] goal (K tac);
       in
-        add_qualified_thm "chain_take" (dname, chain_take_thm) thy
+        add_qualified_simp_thm "chain_take" (dname, thm) thy
       end;
     val (chain_take_thms, thy) =
       fold_map prove_chain_take (take_consts ~~ dnames) thy;
@@ -342,17 +349,17 @@
           (conjuncts dnames deflation_take_thm)) thy;
 
     (* prove strictness of take functions *)
-    fun prove_take_strict (take_const, dname) thy =
+    fun prove_take_strict (deflation_take, dname) thy =
       let
-        val goal = mk_trp (mk_strict (take_const $ Free ("n", natT)));
-        val tac = rtac @{thm deflation_strict} 1
-                  THEN resolve_tac deflation_take_thms 1;
-        val take_strict_thm = Goal.prove_global thy [] [] goal (K tac);
+        val take_strict_thm =
+            Drule.export_without_context
+            (@{thm deflation_strict} OF [deflation_take]);
       in
         add_qualified_thm "take_strict" (dname, take_strict_thm) thy
       end;
     val (take_strict_thms, thy) =
-      fold_map prove_take_strict (take_consts ~~ dnames) thy;
+      fold_map prove_take_strict
+        (deflation_take_thms ~~ dnames) thy;
 
     (* prove take/take rules *)
     fun prove_take_take ((chain_take, deflation_take), dname) thy =
@@ -367,6 +374,19 @@
       fold_map prove_take_take
         (chain_take_thms ~~ deflation_take_thms ~~ dnames) thy;
 
+    (* prove take_below rules *)
+    fun prove_take_below (deflation_take, dname) thy =
+      let
+        val take_below_thm =
+            Drule.export_without_context
+            (@{thm deflation.below} OF [deflation_take]);
+      in
+        add_qualified_thm "take_below" (dname, take_below_thm) thy
+      end;
+    val (take_below_thms, thy) =
+      fold_map prove_take_below
+        (deflation_take_thms ~~ dnames) thy;
+
     (* define finiteness predicates *)
     fun define_finite_const ((tbind, take_const), (lhsT, rhsT)) thy =
       let
--- a/src/HOLCF/Tools/Domain/domain_theorems.ML	Thu Mar 04 22:48:50 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_theorems.ML	Thu Mar 04 18:18:52 2010 -0800
@@ -196,15 +196,20 @@
         pat_rews @ dist_les @ dist_eqs)
 end; (* let *)
 
-fun comp_theorems (comp_dnam, eqs: eq list) thy =
+fun prove_coinduction
+    (comp_dnam, eqs : eq list)
+    (take_lemmas : thm list)
+    (thy : theory) : thm * theory =
 let
-val map_tab = Domain_Take_Proofs.get_map_tab thy;
 
 val dnames = map (fst o fst) eqs;
-val conss  = map  snd        eqs;
 val comp_dname = Sign.full_bname thy comp_dnam;
+fun dc_take dn = %%:(dn^"_take");
+val x_name = idx_name dnames "x"; 
+val n_eqs = length eqs;
 
-val _ = message ("Proving induction properties of domain "^comp_dname^" ...");
+val take_rews =
+    maps (fn dn => PureThy.get_thms thy (dn ^ ".take_rews")) dnames;
 
 (* ----- define bisimulation predicate -------------------------------------- *)
 
@@ -280,6 +285,74 @@
         ||> Sign.parent_path;
 end; (* local *)
 
+(* ----- theorem concerning coinduction ------------------------------------- *)
+
+local
+  val pg = pg' thy;
+  val xs = mapn (fn n => K (x_name n)) 1 dnames;
+  fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1);
+  val take_ss = HOL_ss addsimps (@{thm Rep_CFun_strict1} :: take_rews);
+  val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")"));
+  val _ = trace " Proving coind_lemma...";
+  val coind_lemma =
+    let
+      fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1;
+      fun mk_eqn n dn =
+        (dc_take dn $ %:"n" ` bnd_arg n 0) ===
+        (dc_take dn $ %:"n" ` bnd_arg n 1);
+      fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t));
+      val goal =
+        mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R",
+          Library.foldr mk_all2 (xs,
+            Library.foldr mk_imp (mapn mk_prj 0 dnames,
+              foldr1 mk_conj (mapn mk_eqn 0 dnames)))));
+      fun x_tacs ctxt n x = [
+        rotate_tac (n+1) 1,
+        etac all2E 1,
+        eres_inst_tac ctxt [(("P", 1), sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1,
+        TRY (safe_tac HOL_cs),
+        REPEAT (CHANGED (asm_simp_tac take_ss 1))];
+      fun tacs ctxt = [
+        rtac impI 1,
+        InductTacs.induct_tac ctxt [[SOME "n"]] 1,
+        simp_tac take_ss 1,
+        safe_tac HOL_cs] @
+        flat (mapn (x_tacs ctxt) 0 xs);
+    in pg [ax_bisim_def] goal tacs end;
+in
+  val _ = trace " Proving coind...";
+  val coind = 
+    let
+      fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'"));
+      fun mk_eqn x = %:x === %:(x^"'");
+      val goal =
+        mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===>
+          Logic.list_implies (mapn mk_prj 0 xs,
+            mk_trp (foldr1 mk_conj (map mk_eqn xs)));
+      val tacs =
+        TRY (safe_tac HOL_cs) ::
+        maps (fn take_lemma => [
+          rtac take_lemma 1,
+          cut_facts_tac [coind_lemma] 1,
+          fast_tac HOL_cs 1])
+        take_lemmas;
+    in pg [] goal (K tacs) end;
+end; (* local *)
+
+in
+  (coind, thy)
+end;
+
+fun comp_theorems (comp_dnam, eqs: eq list) thy =
+let
+val map_tab = Domain_Take_Proofs.get_map_tab thy;
+
+val dnames = map (fst o fst) eqs;
+val conss  = map  snd        eqs;
+val comp_dname = Sign.full_bname thy comp_dnam;
+
+val _ = message ("Proving induction properties of domain "^comp_dname^" ...");
+
 val pg = pg' thy;
 
 (* ----- getting the composite axiom and definitions ------------------------ *)
@@ -556,58 +629,7 @@
 
 end; (* local *)
 
-(* ----- theorem concerning coinduction ------------------------------------- *)
-
-local
-  val xs = mapn (fn n => K (x_name n)) 1 dnames;
-  fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1);
-  val take_ss = HOL_ss addsimps (@{thm Rep_CFun_strict1} :: take_rews);
-  val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")"));
-  val _ = trace " Proving coind_lemma...";
-  val coind_lemma =
-    let
-      fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1;
-      fun mk_eqn n dn =
-        (dc_take dn $ %:"n" ` bnd_arg n 0) ===
-        (dc_take dn $ %:"n" ` bnd_arg n 1);
-      fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t));
-      val goal =
-        mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R",
-          Library.foldr mk_all2 (xs,
-            Library.foldr mk_imp (mapn mk_prj 0 dnames,
-              foldr1 mk_conj (mapn mk_eqn 0 dnames)))));
-      fun x_tacs ctxt n x = [
-        rotate_tac (n+1) 1,
-        etac all2E 1,
-        eres_inst_tac ctxt [(("P", 1), sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1,
-        TRY (safe_tac HOL_cs),
-        REPEAT (CHANGED (asm_simp_tac take_ss 1))];
-      fun tacs ctxt = [
-        rtac impI 1,
-        InductTacs.induct_tac ctxt [[SOME "n"]] 1,
-        simp_tac take_ss 1,
-        safe_tac HOL_cs] @
-        flat (mapn (x_tacs ctxt) 0 xs);
-    in pg [ax_bisim_def] goal tacs end;
-in
-  val _ = trace " Proving coind...";
-  val coind = 
-    let
-      fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'"));
-      fun mk_eqn x = %:x === %:(x^"'");
-      val goal =
-        mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===>
-          Logic.list_implies (mapn mk_prj 0 xs,
-            mk_trp (foldr1 mk_conj (map mk_eqn xs)));
-      val tacs =
-        TRY (safe_tac HOL_cs) ::
-        maps (fn take_lemma => [
-          rtac take_lemma 1,
-          cut_facts_tac [coind_lemma] 1,
-          fast_tac HOL_cs 1])
-        take_lemmas;
-    in pg [] goal (K tacs) end;
-end; (* local *)
+val (coind, thy) = prove_coinduction (comp_dnam, eqs) take_lemmas thy;
 
 val inducts = Project_Rule.projections (ProofContext.init thy) ind;
 fun ind_rule (dname, rule) = ((Binding.empty, [rule]), [Induct.induct_type dname]);