src/ZF/ind-syntax.ML
changeset 0 a5a9c433f639
child 6 8ce8c4d13d4d
equal deleted inserted replaced
-1:000000000000 0:a5a9c433f639
       
     1 (*  Title: 	ZF/ind-syntax.ML
       
     2     ID:         $Id$
       
     3     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1993  University of Cambridge
       
     5 
       
     6 Abstract Syntax functions for Inductive Definitions
       
     7 *)
       
     8 
       
     9 
       
    10 (*SHOULD BE ABLE TO DELETE THESE!*)
       
    11 fun flatten_typ sign T = 
       
    12     let val {syn,...} = Sign.rep_sg sign 
       
    13     in  Pretty.str_of (Syntax.pretty_typ syn T)
       
    14     end;
       
    15 fun flatten_term sign t = Pretty.str_of (Sign.pretty_term sign t);
       
    16 
       
    17 (*Add constants to a theory*)
       
    18 infix addconsts;
       
    19 fun thy addconsts const_decs = 
       
    20     extend_theory thy (space_implode "_" (flat (map #1 const_decs)) 
       
    21 		       ^ "_Theory")
       
    22 		  ([], [], [], [], const_decs, None) [];
       
    23 
       
    24 
       
    25 (*Make a definition, lhs==rhs, checking that vars on lhs contain *)
       
    26 fun mk_defpair sign (lhs,rhs) = 
       
    27   let val Const(name,_) = head_of lhs
       
    28       val dummy = assert (term_vars rhs subset term_vars lhs
       
    29 		       andalso
       
    30 		       term_frees rhs subset term_frees lhs
       
    31 		       andalso
       
    32 		       term_tvars rhs subset term_tvars lhs
       
    33 		       andalso
       
    34 		       term_tfrees rhs subset term_tfrees lhs)
       
    35 	          ("Extra variables on RHS in definition of " ^ name)
       
    36   in  (name ^ "_def",
       
    37        flatten_term sign (Logic.mk_equals (lhs,rhs)))
       
    38   end;
       
    39 
       
    40 (*export to Pure/sign?  Used in Provers/simp.ML...*)
       
    41 fun lookup_const sign a = Symtab.lookup(#const_tab (Sign.rep_sg sign), a);
       
    42 
       
    43 (*export to Pure/library?  *)
       
    44 fun assert_all pred l msg_fn = 
       
    45   let fun asl [] = ()
       
    46 	| asl (x::xs) = if pred x then asl xs
       
    47 	                else error (msg_fn x)
       
    48   in  asl l  end;
       
    49 
       
    50 
       
    51 (** Abstract syntax definitions for FOL and ZF **)
       
    52 
       
    53 val iT = Type("i",[])
       
    54 and oT = Type("o",[]);
       
    55 
       
    56 fun ap t u = t$u;
       
    57 fun app t (u1,u2) = t $ u1 $ u2;
       
    58 
       
    59 (*Given u expecting arguments of types [T1,...,Tn], create term of 
       
    60   type T1*...*Tn => i using split*)
       
    61 fun ap_split split u [ ]   = Abs("null", iT, u)
       
    62   | ap_split split u [_]   = u
       
    63   | ap_split split u [_,_] = split $ u
       
    64   | ap_split split u (T::Ts) = 
       
    65       split $ (Abs("v", T, ap_split split (u $ Bound(length Ts - 2)) Ts));
       
    66 
       
    67 val conj = Const("op &", [oT,oT]--->oT)
       
    68 and disj = Const("op |", [oT,oT]--->oT)
       
    69 and imp = Const("op -->", [oT,oT]--->oT);
       
    70 
       
    71 val eq_const = Const("op =", [iT,iT]--->oT);
       
    72 
       
    73 val mem_const = Const("op :", [iT,iT]--->oT);
       
    74 
       
    75 val exists_const = Const("Ex", [iT-->oT]--->oT);
       
    76 fun mk_exists (Free(x,T),P) = exists_const $ (absfree (x,T,P));
       
    77 
       
    78 val all_const = Const("All", [iT-->oT]--->oT);
       
    79 fun mk_all (Free(x,T),P) = all_const $ (absfree (x,T,P));
       
    80 
       
    81 (*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *)
       
    82 fun mk_all_imp (A,P) = 
       
    83     all_const $ Abs("v", iT, imp $ (mem_const $ Bound 0 $ A) $ (P $ Bound 0));
       
    84 
       
    85 
       
    86 val Part_const = Const("Part", [iT,iT-->iT]--->iT);
       
    87 
       
    88 val Collect_const = Const("Collect", [iT,iT-->oT]--->iT);
       
    89 fun mk_Collect (a,D,t) = Collect_const $ D $ absfree(a, iT, t);
       
    90 
       
    91 val Trueprop = Const("Trueprop",oT-->propT);
       
    92 fun mk_tprop P = Trueprop $ P;
       
    93 fun dest_tprop (Const("Trueprop",_) $ P) = P;
       
    94 
       
    95 (*** Tactic for folding constructor definitions ***)
       
    96 
       
    97 (*The depth of injections in a constructor function*)
       
    98 fun inject_depth (Const _ $ t) = 1 + inject_depth t
       
    99   | inject_depth t             = 0;
       
   100 
       
   101 val rhs_of_thm = #2 o Logic.dest_equals o #prop o rep_thm;
       
   102 
       
   103 (*There are critical pairs!  E.g. K == Inl(0),  S == Inr(Inl(0))
       
   104   Folds longest definitions first to avoid folding subexpressions of an rhs.*)
       
   105 fun fold_con_tac defs =
       
   106   let val keylist = make_keylist (inject_depth o rhs_of_thm) defs;
       
   107       val keys = distinct (sort op> (map #2 keylist));
       
   108       val deflists = map (keyfilter keylist) keys
       
   109   in  EVERY (map fold_tac deflists)  end;
       
   110 
       
   111 (*Prove a goal stated as a term, with exception handling*)
       
   112 fun prove_term sign defs (P,tacsf) = 
       
   113     let val ct = Sign.cterm_of sign P
       
   114     in  prove_goalw_cterm defs ct tacsf
       
   115 	handle e => (writeln ("Exception in proof of\n" ^
       
   116 			       Sign.string_of_cterm ct); 
       
   117 		     raise e)
       
   118     end;
       
   119 
       
   120 (*Read an assumption in the given theory*)
       
   121 fun assume_read thy a = assume (Sign.read_cterm (sign_of thy) (a,propT));
       
   122 
       
   123 (*Make distinct individual variables a1, a2, a3, ..., an. *)
       
   124 fun mk_frees a [] = []
       
   125   | mk_frees a (T::Ts) = Free(a,T) :: mk_frees (bump_string a) Ts;
       
   126 
       
   127 (*Used by intr-elim.ML and in individual datatype definitions*)
       
   128 val basic_monos = [subset_refl, imp_refl, disj_mono, conj_mono, 
       
   129 		   ex_mono, Collect_mono, Part_mono, in_mono];
       
   130 
       
   131 fun rule_concl rl = 
       
   132   let val Const("op :",_) $ t $ X = dest_tprop (Logic.strip_imp_concl rl)
       
   133   in (t,X) end
       
   134   handle _ => error "Conclusion of rule should be a set membership";
       
   135 
       
   136 (*For deriving cases rules.  CollectD2 discards the domain, which is redundant;
       
   137   read_instantiate replaces a propositional variable by a formula variable*)
       
   138 val equals_CollectD = 
       
   139     read_instantiate [("W","?Q")]
       
   140         (make_elim (equalityD1 RS subsetD RS CollectD2));
       
   141 
       
   142 
       
   143 (*From HOL/ex/meson.ML: raises exception if no rules apply -- unlike RL*)
       
   144 fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls))
       
   145   | tryres (th, []) = raise THM("tryres", 0, [th]);
       
   146 
       
   147 fun gen_make_elim elim_rls rl = 
       
   148       standard (tryres (rl, elim_rls @ [revcut_rl]));
       
   149 
       
   150 (** For constructor.ML **)
       
   151 
       
   152 (*Avoids duplicate definitions by removing constants already declared mixfix*)
       
   153 fun remove_mixfixes None decs = decs
       
   154   | remove_mixfixes (Some sext) decs =
       
   155       let val mixtab = Symtab.st_of_declist(Syntax.constants sext, Symtab.null)
       
   156 	  fun is_mix c = case Symtab.lookup(mixtab,c) of
       
   157 			     None=>false | Some _ => true
       
   158       in  map (fn (cs,styp)=> (filter_out is_mix cs, styp)) decs
       
   159       end;
       
   160 
       
   161 fun ext_constants None        = []
       
   162   | ext_constants (Some sext) = Syntax.constants sext;
       
   163 
       
   164 
       
   165 (*Could go to FOL, but it's hardly general*)
       
   166 val [def] = goal IFOL.thy "a==b ==> a=c <-> c=b";
       
   167 by (rewtac def);
       
   168 by (rtac iffI 1);
       
   169 by (REPEAT (etac sym 1));
       
   170 val def_swap_iff = result();
       
   171 
       
   172 val def_trans = prove_goal IFOL.thy "[| f==g;  g(a)=b |] ==> f(a)=b"
       
   173   (fn [rew,prem] => [ rewtac rew, rtac prem 1 ]);
       
   174 
       
   175