src/HOL/ex/Simult.ML
changeset 969 b051e2fc2e34
child 1266 3ae9fe3c0f68
equal deleted inserted replaced
968:3cdaa8724175 969:b051e2fc2e34
       
     1 (*  Title: 	HOL/ex/Simult.ML
       
     2     ID:         $Id$
       
     3     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1993  University of Cambridge
       
     5 
       
     6 Primitives for simultaneous recursive type definitions
       
     7   includes worked example of trees & forests
       
     8 
       
     9 This is essentially the same data structure that on ex/term.ML, which is
       
    10 simpler because it uses list as a new type former.  The approach in this
       
    11 file may be superior for other simultaneous recursions.
       
    12 *)
       
    13 
       
    14 open Simult;
       
    15 
       
    16 (*** Monotonicity and unfolding of the function ***)
       
    17 
       
    18 goal Simult.thy "mono(%Z.  A <*> Part Z In1 \
       
    19 \                      <+> ({Numb(0)} <+> Part Z In0 <*> Part Z In1))";
       
    20 by (REPEAT (ares_tac [monoI, subset_refl, usum_mono, uprod_mono,
       
    21 		      Part_mono] 1));
       
    22 qed "TF_fun_mono";
       
    23 
       
    24 val TF_unfold = TF_fun_mono RS (TF_def RS def_lfp_Tarski);
       
    25 
       
    26 goalw Simult.thy [TF_def] "!!A B. A<=B ==> TF(A) <= TF(B)";
       
    27 by (REPEAT (ares_tac [lfp_mono, subset_refl, usum_mono, uprod_mono] 1));
       
    28 qed "TF_mono";
       
    29 
       
    30 goalw Simult.thy [TF_def] "TF(sexp) <= sexp";
       
    31 by (rtac lfp_lowerbound 1);
       
    32 by (fast_tac (univ_cs addIs  sexp.intrs@[sexp_In0I, sexp_In1I]
       
    33                       addSEs [PartE]) 1);
       
    34 qed "TF_sexp";
       
    35 
       
    36 (* A <= sexp ==> TF(A) <= sexp *)
       
    37 val TF_subset_sexp = standard
       
    38     (TF_mono RS (TF_sexp RSN (2,subset_trans)));
       
    39 
       
    40 
       
    41 (** Elimination -- structural induction on the set TF **)
       
    42 
       
    43 val TF_Rep_defs = [TCONS_def,FNIL_def,FCONS_def,NIL_def,CONS_def];
       
    44 
       
    45 val major::prems = goalw Simult.thy TF_Rep_defs
       
    46  "[| i: TF(A);  \
       
    47 \    !!M N. [| M: A;  N: Part (TF A) In1;  R(N) |] ==> R(TCONS M N);	\
       
    48 \    R(FNIL);        		\
       
    49 \    !!M N. [| M:  Part (TF A) In0;  N: Part (TF A) In1;  R(M);  R(N) \
       
    50 \            |] ==> R(FCONS M N)    \
       
    51 \    |] ==> R(i)";
       
    52 by (rtac ([TF_def, TF_fun_mono, major] MRS def_induct) 1);
       
    53 by (fast_tac (set_cs addIs (prems@[PartI])
       
    54 		       addEs [usumE, uprodE, PartE]) 1);
       
    55 qed "TF_induct";
       
    56 
       
    57 (*This lemma replaces a use of subgoal_tac to prove tree_forest_induct*)
       
    58 val prems = goalw Simult.thy [Part_def]
       
    59  "! M: TF(A). (M: Part (TF A) In0 --> P(M)) & (M: Part (TF A) In1 --> Q(M)) \
       
    60 \ ==> (! M: Part (TF A) In0. P(M)) & (! M: Part (TF A) In1. Q(M))";
       
    61 by (cfast_tac prems 1);
       
    62 qed "TF_induct_lemma";
       
    63 
       
    64 val uplus_cs = set_cs addSIs [PartI]
       
    65 		      addSDs [In0_inject, In1_inject]
       
    66 		      addSEs [In0_neq_In1, In1_neq_In0, PartE];
       
    67 
       
    68 (*Could prove  ~ TCONS M N : Part (TF A) In1  etc. *)
       
    69 
       
    70 (*Induction on TF with separate predicates P, Q*)
       
    71 val prems = goalw Simult.thy TF_Rep_defs
       
    72     "[| !!M N. [| M: A;  N: Part (TF A) In1;  Q(N) |] ==> P(TCONS M N); \
       
    73 \       Q(FNIL);        \
       
    74 \       !!M N. [| M:  Part (TF A) In0;  N: Part (TF A) In1;  P(M);  Q(N) \
       
    75 \               |] ==> Q(FCONS M N)     \
       
    76 \    |] ==> (! M: Part (TF A) In0. P(M)) & (! N: Part (TF A) In1. Q(N))";
       
    77 by (rtac (ballI RS TF_induct_lemma) 1);
       
    78 by (etac TF_induct 1);
       
    79 by (rewrite_goals_tac TF_Rep_defs);
       
    80 by (ALLGOALS (fast_tac (uplus_cs addIs prems)));
       
    81 (*29 secs??*)
       
    82 qed "Tree_Forest_induct";
       
    83 
       
    84 (*Induction for the abstract types 'a tree, 'a forest*)
       
    85 val prems = goalw Simult.thy [Tcons_def,Fnil_def,Fcons_def]
       
    86     "[| !!x ts. Q(ts) ==> P(Tcons x ts);     \
       
    87 \	Q(Fnil);        \
       
    88 \       !!t ts. [| P(t);  Q(ts) |] ==> Q(Fcons t ts)    \
       
    89 \    |] ==> (! t. P(t)) & (! ts. Q(ts))";
       
    90 by (res_inst_tac [("P1","%z.P(Abs_Tree(z))"),
       
    91 		  ("Q1","%z.Q(Abs_Forest(z))")] 
       
    92     (Tree_Forest_induct RS conjE) 1);
       
    93 (*Instantiates ?A1 to range(Leaf). *)
       
    94 by (fast_tac (set_cs addSEs [Rep_Tree_inverse RS subst, 
       
    95 			     Rep_Forest_inverse RS subst] 
       
    96 	             addSIs [Rep_Tree,Rep_Forest]) 4);
       
    97 (*Cannot use simplifier: the rewrites work in the wrong direction!*)
       
    98 by (ALLGOALS (fast_tac (set_cs addSEs [Abs_Tree_inverse RS subst,
       
    99                           Abs_Forest_inverse RS subst] 
       
   100 	             addSIs prems)));
       
   101 qed "tree_forest_induct";
       
   102 
       
   103 
       
   104 
       
   105 (*** Isomorphisms ***)
       
   106 
       
   107 goal Simult.thy "inj(Rep_Tree)";
       
   108 by (rtac inj_inverseI 1);
       
   109 by (rtac Rep_Tree_inverse 1);
       
   110 qed "inj_Rep_Tree";
       
   111 
       
   112 goal Simult.thy "inj_onto Abs_Tree (Part (TF(range Leaf)) In0)";
       
   113 by (rtac inj_onto_inverseI 1);
       
   114 by (etac Abs_Tree_inverse 1);
       
   115 qed "inj_onto_Abs_Tree";
       
   116 
       
   117 goal Simult.thy "inj(Rep_Forest)";
       
   118 by (rtac inj_inverseI 1);
       
   119 by (rtac Rep_Forest_inverse 1);
       
   120 qed "inj_Rep_Forest";
       
   121 
       
   122 goal Simult.thy "inj_onto Abs_Forest (Part (TF(range Leaf)) In1)";
       
   123 by (rtac inj_onto_inverseI 1);
       
   124 by (etac Abs_Forest_inverse 1);
       
   125 qed "inj_onto_Abs_Forest";
       
   126 
       
   127 (** Introduction rules for constructors **)
       
   128 
       
   129 (* c : A <*> Part (TF A) In1 
       
   130         <+> {Numb(0)} <+> Part (TF A) In0 <*> Part (TF A) In1 ==> c : TF(A) *)
       
   131 val TF_I = TF_unfold RS equalityD2 RS subsetD;
       
   132 
       
   133 (*For reasoning about the representation*)
       
   134 val TF_Rep_cs = uplus_cs addIs [TF_I, uprodI, usum_In0I, usum_In1I]
       
   135 	                 addSEs [Scons_inject];
       
   136 
       
   137 val prems = goalw Simult.thy TF_Rep_defs
       
   138     "[| a: A;  M: Part (TF A) In1 |] ==> TCONS a M : Part (TF A) In0";
       
   139 by (fast_tac (TF_Rep_cs addIs prems) 1);
       
   140 qed "TCONS_I";
       
   141 
       
   142 (* FNIL is a TF(A) -- this also justifies the type definition*)
       
   143 goalw Simult.thy TF_Rep_defs "FNIL: Part (TF A) In1";
       
   144 by (fast_tac TF_Rep_cs 1);
       
   145 qed "FNIL_I";
       
   146 
       
   147 val prems = goalw Simult.thy TF_Rep_defs
       
   148     "[| M: Part (TF A) In0;  N: Part (TF A) In1 |] ==> \
       
   149 \    FCONS M N : Part (TF A) In1";
       
   150 by (fast_tac (TF_Rep_cs addIs prems) 1);
       
   151 qed "FCONS_I";
       
   152 
       
   153 (** Injectiveness of TCONS and FCONS **)
       
   154 
       
   155 goalw Simult.thy TF_Rep_defs "(TCONS K M=TCONS L N) = (K=L & M=N)";
       
   156 by (fast_tac TF_Rep_cs 1);
       
   157 qed "TCONS_TCONS_eq";
       
   158 bind_thm ("TCONS_inject", (TCONS_TCONS_eq RS iffD1 RS conjE));
       
   159 
       
   160 goalw Simult.thy TF_Rep_defs "(FCONS K M=FCONS L N) = (K=L & M=N)";
       
   161 by (fast_tac TF_Rep_cs 1);
       
   162 qed "FCONS_FCONS_eq";
       
   163 bind_thm ("FCONS_inject", (FCONS_FCONS_eq RS iffD1 RS conjE));
       
   164 
       
   165 (** Distinctness of TCONS, FNIL and FCONS **)
       
   166 
       
   167 goalw Simult.thy TF_Rep_defs "TCONS M N ~= FNIL";
       
   168 by (fast_tac TF_Rep_cs 1);
       
   169 qed "TCONS_not_FNIL";
       
   170 bind_thm ("FNIL_not_TCONS", (TCONS_not_FNIL RS not_sym));
       
   171 
       
   172 bind_thm ("TCONS_neq_FNIL", (TCONS_not_FNIL RS notE));
       
   173 val FNIL_neq_TCONS = sym RS TCONS_neq_FNIL;
       
   174 
       
   175 goalw Simult.thy TF_Rep_defs "FCONS M N ~= FNIL";
       
   176 by (fast_tac TF_Rep_cs 1);
       
   177 qed "FCONS_not_FNIL";
       
   178 bind_thm ("FNIL_not_FCONS", (FCONS_not_FNIL RS not_sym));
       
   179 
       
   180 bind_thm ("FCONS_neq_FNIL", (FCONS_not_FNIL RS notE));
       
   181 val FNIL_neq_FCONS = sym RS FCONS_neq_FNIL;
       
   182 
       
   183 goalw Simult.thy TF_Rep_defs "TCONS M N ~= FCONS K L";
       
   184 by (fast_tac TF_Rep_cs 1);
       
   185 qed "TCONS_not_FCONS";
       
   186 bind_thm ("FCONS_not_TCONS", (TCONS_not_FCONS RS not_sym));
       
   187 
       
   188 bind_thm ("TCONS_neq_FCONS", (TCONS_not_FCONS RS notE));
       
   189 val FCONS_neq_TCONS = sym RS TCONS_neq_FCONS;
       
   190 
       
   191 (*???? Too many derived rules ????
       
   192   Automatically generate symmetric forms?  Always expand TF_Rep_defs? *)
       
   193 
       
   194 (** Injectiveness of Tcons and Fcons **)
       
   195 
       
   196 (*For reasoning about abstract constructors*)
       
   197 val TF_cs = set_cs addSIs [Rep_Tree, Rep_Forest, TCONS_I, FNIL_I, FCONS_I]
       
   198 	           addSEs [TCONS_inject, FCONS_inject,
       
   199 			   TCONS_neq_FNIL, FNIL_neq_TCONS,
       
   200 			   FCONS_neq_FNIL, FNIL_neq_FCONS,
       
   201 			   TCONS_neq_FCONS, FCONS_neq_TCONS]
       
   202 		   addSDs [inj_onto_Abs_Tree RS inj_ontoD,
       
   203 			   inj_onto_Abs_Forest RS inj_ontoD,
       
   204 			   inj_Rep_Tree RS injD, inj_Rep_Forest RS injD,
       
   205 			   Leaf_inject];
       
   206 
       
   207 goalw Simult.thy [Tcons_def] "(Tcons x xs=Tcons y ys) = (x=y & xs=ys)";
       
   208 by (fast_tac TF_cs 1);
       
   209 qed "Tcons_Tcons_eq";
       
   210 bind_thm ("Tcons_inject", (Tcons_Tcons_eq RS iffD1 RS conjE));
       
   211 
       
   212 goalw Simult.thy [Fcons_def,Fnil_def] "Fcons x xs ~= Fnil";
       
   213 by (fast_tac TF_cs 1);
       
   214 qed "Fcons_not_Fnil";
       
   215 
       
   216 bind_thm ("Fcons_neq_Fnil", Fcons_not_Fnil RS notE);
       
   217 val Fnil_neq_Fcons = sym RS Fcons_neq_Fnil;
       
   218 
       
   219 
       
   220 (** Injectiveness of Fcons **)
       
   221 
       
   222 goalw Simult.thy [Fcons_def] "(Fcons x xs=Fcons y ys) = (x=y & xs=ys)";
       
   223 by (fast_tac TF_cs 1);
       
   224 qed "Fcons_Fcons_eq";
       
   225 bind_thm ("Fcons_inject", Fcons_Fcons_eq RS iffD1 RS conjE);
       
   226 
       
   227 
       
   228 (*** TF_rec -- by wf recursion on pred_sexp ***)
       
   229 
       
   230 val TF_rec_unfold =
       
   231     wf_pred_sexp RS wf_trancl RS (TF_rec_def RS def_wfrec);
       
   232 
       
   233 (** conversion rules for TF_rec **)
       
   234 
       
   235 goalw Simult.thy [TCONS_def]
       
   236     "!!M N. [| M: sexp;  N: sexp |] ==> 	\
       
   237 \           TF_rec (TCONS M N) b c d = b M N (TF_rec N b c d)";
       
   238 by (rtac (TF_rec_unfold RS trans) 1);
       
   239 by (simp_tac (HOL_ss addsimps [Case_In0, Split]) 1);
       
   240 by (asm_simp_tac (pred_sexp_ss addsimps [In0_def]) 1);
       
   241 qed "TF_rec_TCONS";
       
   242 
       
   243 goalw Simult.thy [FNIL_def] "TF_rec FNIL b c d = c";
       
   244 by (rtac (TF_rec_unfold RS trans) 1);
       
   245 by (simp_tac (HOL_ss addsimps [Case_In1, List_case_NIL]) 1);
       
   246 qed "TF_rec_FNIL";
       
   247 
       
   248 goalw Simult.thy [FCONS_def]
       
   249  "!!M N. [| M: sexp;  N: sexp |] ==> 	\
       
   250 \        TF_rec (FCONS M N) b c d = d M N (TF_rec M b c d) (TF_rec N b c d)";
       
   251 by (rtac (TF_rec_unfold RS trans) 1);
       
   252 by (simp_tac (HOL_ss addsimps [Case_In1, List_case_CONS]) 1);
       
   253 by (asm_simp_tac (pred_sexp_ss addsimps [CONS_def,In1_def]) 1);
       
   254 qed "TF_rec_FCONS";
       
   255 
       
   256 
       
   257 (*** tree_rec, forest_rec -- by TF_rec ***)
       
   258 
       
   259 val Rep_Tree_in_sexp =
       
   260     [range_Leaf_subset_sexp RS TF_subset_sexp RS (Part_subset RS subset_trans),
       
   261      Rep_Tree] MRS subsetD;
       
   262 val Rep_Forest_in_sexp =
       
   263     [range_Leaf_subset_sexp RS TF_subset_sexp RS (Part_subset RS subset_trans),
       
   264      Rep_Forest] MRS subsetD;
       
   265 
       
   266 val tf_rec_simps = [TF_rec_TCONS, TF_rec_FNIL, TF_rec_FCONS,
       
   267 		    TCONS_I, FNIL_I, FCONS_I, Rep_Tree, Rep_Forest,
       
   268 		    Rep_Tree_inverse, Rep_Forest_inverse,
       
   269 		    Abs_Tree_inverse, Abs_Forest_inverse,
       
   270 		    inj_Leaf, Inv_f_f, sexp.LeafI, range_eqI,
       
   271 		    Rep_Tree_in_sexp, Rep_Forest_in_sexp];
       
   272 val tf_rec_ss = HOL_ss addsimps tf_rec_simps;
       
   273 
       
   274 goalw Simult.thy [tree_rec_def, forest_rec_def, Tcons_def]
       
   275     "tree_rec (Tcons a tf) b c d = b a tf (forest_rec tf b c d)";
       
   276 by (simp_tac tf_rec_ss 1);
       
   277 qed "tree_rec_Tcons";
       
   278 
       
   279 goalw Simult.thy [forest_rec_def, Fnil_def] "forest_rec Fnil b c d = c";
       
   280 by (simp_tac tf_rec_ss 1);
       
   281 qed "forest_rec_Fnil";
       
   282 
       
   283 goalw Simult.thy [tree_rec_def, forest_rec_def, Fcons_def]
       
   284     "forest_rec (Fcons t tf) b c d = \
       
   285 \    d t tf (tree_rec t b c d) (forest_rec tf b c d)";
       
   286 by (simp_tac tf_rec_ss 1);
       
   287 qed "forest_rec_Cons";