src/HOL/ex/Term.ML
changeset 3125 3f0ab2c306f7
parent 3124 1c0dfa7ebb72
child 3126 feb7a5d01c1e
equal deleted inserted replaced
3124:1c0dfa7ebb72 3125:3f0ab2c306f7
     1 (*  Title:      HOL/ex/Term
       
     2     ID:         $Id$
       
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1992  University of Cambridge
       
     5 
       
     6 Terms over a given alphabet -- function applications; illustrates list functor
       
     7   (essentially the same type as in Trees & Forests)
       
     8 *)
       
     9 
       
    10 open Term;
       
    11 
       
    12 (*** Monotonicity and unfolding of the function ***)
       
    13 
       
    14 goal Term.thy "term(A) = A <*> list(term(A))";
       
    15 by (fast_tac (!claset addSIs (equalityI :: term.intrs)
       
    16                       addEs [term.elim]) 1);
       
    17 qed "term_unfold";
       
    18 
       
    19 (*This justifies using term in other recursive type definitions*)
       
    20 goalw Term.thy term.defs "!!A B. A<=B ==> term(A) <= term(B)";
       
    21 by (REPEAT (ares_tac ([lfp_mono, list_mono] @ basic_monos) 1));
       
    22 qed "term_mono";
       
    23 
       
    24 (** Type checking -- term creates well-founded sets **)
       
    25 
       
    26 goalw Term.thy term.defs "term(sexp) <= sexp";
       
    27 by (rtac lfp_lowerbound 1);
       
    28 by (fast_tac (!claset addIs [sexp.SconsI, list_sexp RS subsetD]) 1);
       
    29 qed "term_sexp";
       
    30 
       
    31 (* A <= sexp ==> term(A) <= sexp *)
       
    32 bind_thm ("term_subset_sexp", ([term_mono, term_sexp] MRS subset_trans));
       
    33 
       
    34 
       
    35 (** Elimination -- structural induction on the set term(A) **)
       
    36 
       
    37 (*Induction for the set term(A) *)
       
    38 val [major,minor] = goal Term.thy 
       
    39     "[| M: term(A);  \
       
    40 \       !!x zs. [| x: A;  zs: list(term(A));  zs: list({x.R(x)}) \
       
    41 \               |] ==> R(x$zs)  \
       
    42 \    |] ==> R(M)";
       
    43 by (rtac (major RS term.induct) 1);
       
    44 by (REPEAT (eresolve_tac ([minor] @
       
    45                 ([Int_lower1,Int_lower2] RL [list_mono RS subsetD])) 1));
       
    46 (*Proof could also use  mono_Int RS subsetD RS IntE *)
       
    47 qed "Term_induct";
       
    48 
       
    49 (*Induction on term(A) followed by induction on list *)
       
    50 val major::prems = goal Term.thy
       
    51     "[| M: term(A);  \
       
    52 \       !!x.      [| x: A |] ==> R(x$NIL);  \
       
    53 \       !!x z zs. [| x: A;  z: term(A);  zs: list(term(A));  R(x$zs)  \
       
    54 \                 |] ==> R(x $ CONS z zs)  \
       
    55 \    |] ==> R(M)";
       
    56 by (rtac (major RS Term_induct) 1);
       
    57 by (etac list.induct 1);
       
    58 by (REPEAT (ares_tac prems 1));
       
    59 qed "Term_induct2";
       
    60 
       
    61 (*** Structural Induction on the abstract type 'a term ***)
       
    62 
       
    63 val Rep_term_in_sexp =
       
    64     Rep_term RS (range_Leaf_subset_sexp RS term_subset_sexp RS subsetD);
       
    65 
       
    66 (*Induction for the abstract type 'a term*)
       
    67 val prems = goalw Term.thy [App_def,Rep_Tlist_def,Abs_Tlist_def]
       
    68     "[| !!x ts. (ALL t: set_of_list ts. R t) ==> R(App x ts)  \
       
    69 \    |] ==> R(t)";
       
    70 by (rtac (Rep_term_inverse RS subst) 1);   (*types force good instantiation*)
       
    71 by (res_inst_tac [("P","Rep_term(t) : sexp")] conjunct2 1);
       
    72 by (rtac (Rep_term RS Term_induct) 1);
       
    73 by (REPEAT (ares_tac [conjI, sexp.SconsI, term_subset_sexp RS 
       
    74     list_subset_sexp, range_Leaf_subset_sexp] 1
       
    75      ORELSE etac rev_subsetD 1));
       
    76 by (eres_inst_tac [("A1","term(?u)"), ("f1","Rep_term"), ("g1","Abs_term")]
       
    77         (Abs_map_inverse RS subst) 1);
       
    78 by (rtac (range_Leaf_subset_sexp RS term_subset_sexp) 1);
       
    79 by (etac Abs_term_inverse 1);
       
    80 by (etac rangeE 1);
       
    81 by (hyp_subst_tac 1);
       
    82 by (resolve_tac prems 1);
       
    83 by (etac list.induct 1);
       
    84 by (etac CollectE 2);
       
    85 by (stac Abs_map_CONS 2);
       
    86 by (etac conjunct1 2);
       
    87 by (etac rev_subsetD 2);
       
    88 by (rtac list_subset_sexp 2);
       
    89 by (ALLGOALS Asm_simp_tac);
       
    90 by (ALLGOALS Fast_tac);
       
    91 qed "term_induct";
       
    92 
       
    93 (*Induction for the abstract type 'a term*)
       
    94 val prems = goal Term.thy 
       
    95     "[| !!x. R(App x Nil);  \
       
    96 \       !!x t ts. R(App x ts) ==> R(App x (t#ts))  \
       
    97 \    |] ==> R(t)";
       
    98 by (rtac term_induct 1);  (*types force good instantiation*)
       
    99 by (etac rev_mp 1);
       
   100 by (rtac list_induct2 1);  (*types force good instantiation*)
       
   101 by (ALLGOALS (asm_simp_tac (!simpset addsimps prems)));
       
   102 qed "term_induct2";
       
   103 
       
   104 (*Perform induction on xs. *)
       
   105 fun term_ind2_tac a i = 
       
   106     EVERY [res_inst_tac [("t",a)] term_induct2 i,
       
   107            rename_last_tac a ["1","s"] (i+1)];
       
   108 
       
   109 
       
   110 
       
   111 (*** Term_rec -- by wf recursion on pred_sexp ***)
       
   112 
       
   113 goal Term.thy
       
   114    "(%M. Term_rec M d) = wfrec (trancl pred_sexp) \
       
   115                              \ (%g. Split(%x y. d x y (Abs_map g y)))";
       
   116 by (simp_tac (HOL_ss addsimps [Term_rec_def]) 1);
       
   117 bind_thm("Term_rec_unfold", (wf_pred_sexp RS wf_trancl) RS 
       
   118                             ((result() RS eq_reflection) RS def_wfrec));
       
   119 
       
   120 (*---------------------------------------------------------------------------
       
   121  * Old:
       
   122  * val Term_rec_unfold =
       
   123  *     wf_pred_sexp RS wf_trancl RS (Term_rec_def RS def_wfrec);
       
   124  *---------------------------------------------------------------------------*)
       
   125 
       
   126 (** conversion rules **)
       
   127 
       
   128 val [prem] = goal Term.thy
       
   129     "N: list(term(A)) ==>  \
       
   130 \    !M. (N,M): pred_sexp^+ --> \
       
   131 \        Abs_map (cut h (pred_sexp^+) M) N = \
       
   132 \        Abs_map h N";
       
   133 by (rtac (prem RS list.induct) 1);
       
   134 by (Simp_tac 1);
       
   135 by (strip_tac 1);
       
   136 by (etac (pred_sexp_CONS_D RS conjE) 1);
       
   137 by (asm_simp_tac (!simpset addsimps [trancl_pred_sexpD1]) 1);
       
   138 qed "Abs_map_lemma";
       
   139 
       
   140 val [prem1,prem2,A_subset_sexp] = goal Term.thy
       
   141     "[| M: sexp;  N: list(term(A));  A<=sexp |] ==> \
       
   142 \    Term_rec (M$N) d = d M N (Abs_map (%Z. Term_rec Z d) N)";
       
   143 by (rtac (Term_rec_unfold RS trans) 1);
       
   144 by (simp_tac (HOL_ss addsimps
       
   145       [Split,
       
   146        prem2 RS Abs_map_lemma RS spec RS mp, pred_sexpI2 RS r_into_trancl,
       
   147        prem1, prem2 RS rev_subsetD, list_subset_sexp,
       
   148        term_subset_sexp, A_subset_sexp]) 1);
       
   149 qed "Term_rec";
       
   150 
       
   151 (*** term_rec -- by Term_rec ***)
       
   152 
       
   153 local
       
   154   val Rep_map_type1 = read_instantiate_sg (sign_of Term.thy)
       
   155                         [("f","Rep_term")] Rep_map_type;
       
   156   val Rep_Tlist = Rep_term RS Rep_map_type1;
       
   157   val Rep_Term_rec = range_Leaf_subset_sexp RSN (2,Rep_Tlist RSN(2,Term_rec));
       
   158 
       
   159   (*Now avoids conditional rewriting with the premise N: list(term(A)),
       
   160     since A will be uninstantiated and will cause rewriting to fail. *)
       
   161   val term_rec_ss = HOL_ss
       
   162       addsimps [Rep_Tlist RS (rangeI RS term.APP_I RS Abs_term_inverse),  
       
   163                 Rep_term_in_sexp, Rep_Term_rec, Rep_term_inverse, inj_Leaf,
       
   164                 inv_f_f, Abs_Rep_map, map_ident2, sexp.LeafI]
       
   165 in
       
   166 
       
   167 val term_rec = prove_goalw Term.thy
       
   168          [term_rec_def, App_def, Rep_Tlist_def, Abs_Tlist_def]
       
   169     "term_rec (App f ts) d = d f ts (map (%t. term_rec t d) ts)"
       
   170  (fn _ => [simp_tac term_rec_ss 1])
       
   171 
       
   172 end;