| author | paulson | 
| Fri, 18 Apr 1997 11:47:11 +0200 | |
| changeset 2981 | aa5aeb6467c6 | 
| parent 1746 | f0c6aabc6c02 | 
| child 4807 | 013ba4c43832 | 
| permissions | -rw-r--r-- | 
| 1465 | 1  | 
(* Title: HOL/ind_syntax.ML  | 
| 923 | 2  | 
ID: $Id$  | 
| 1465 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 923 | 4  | 
Copyright 1994 University of Cambridge  | 
5  | 
||
6  | 
Abstract Syntax functions for Inductive Definitions  | 
|
7  | 
See also hologic.ML and ../Pure/section-utils.ML  | 
|
8  | 
*)  | 
|
9  | 
||
10  | 
(*The structure protects these items from redeclaration (somewhat!). The  | 
|
11  | 
datatype definitions in theory files refer to these items by name!  | 
|
12  | 
*)  | 
|
13  | 
structure Ind_Syntax =  | 
|
14  | 
struct  | 
|
15  | 
||
16  | 
(** Abstract syntax definitions for HOL **)  | 
|
17  | 
||
18  | 
open HOLogic;  | 
|
19  | 
||
20  | 
fun Int_const T =  | 
|
21  | 
let val sT = mk_setT T  | 
|
22  | 
  in  Const("op Int", [sT,sT]--->sT)  end;
 | 
|
23  | 
||
24  | 
fun mk_exists (Free(x,T),P) = exists_const T $ (absfree (x,T,P));  | 
|
25  | 
||
26  | 
fun mk_all (Free(x,T),P) = all_const T $ (absfree (x,T,P));  | 
|
27  | 
||
28  | 
(*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *)  | 
|
29  | 
fun mk_all_imp (A,P) =  | 
|
30  | 
let val T = dest_setT (fastype_of A)  | 
|
31  | 
  in  all_const T $ Abs("v", T, imp $ (mk_mem (Bound 0, A)) $ (P $ Bound 0))
 | 
|
32  | 
end;  | 
|
33  | 
||
34  | 
(** Disjoint sum type **)  | 
|
35  | 
||
36  | 
fun mk_sum (T1,T2) = Type("+", [T1,T2]);
 | 
|
| 1465 | 37  | 
val Inl = Const("Inl", dummyT)
 | 
38  | 
and Inr = Const("Inr", dummyT);         (*correct types added later!*)
 | 
|
39  | 
(*val elim      = Const("case", [iT-->iT, iT-->iT, iT]--->iT)*)
 | 
|
| 923 | 40  | 
|
41  | 
fun summands (Type("+", [T1,T2])) = summands T1 @ summands T2
 | 
|
42  | 
| summands T = [T];  | 
|
43  | 
||
44  | 
(*Given the destination type, fills in correct types of an Inl/Inr nest*)  | 
|
45  | 
fun mend_sum_types (h,T) =  | 
|
46  | 
(case (h,T) of  | 
|
| 1465 | 47  | 
         (Const("Inl",_) $ h1, Type("+", [T1,T2])) =>
 | 
48  | 
             Const("Inl", T1 --> T) $ (mend_sum_types (h1, T1))
 | 
|
| 923 | 49  | 
       | (Const("Inr",_) $ h2, Type("+", [T1,T2])) =>
 | 
| 1465 | 50  | 
             Const("Inr", T2 --> T) $ (mend_sum_types (h2, T2))
 | 
| 923 | 51  | 
| _ => h);  | 
52  | 
||
53  | 
||
54  | 
||
55  | 
(*simple error-checking in the premises of an inductive definition*)  | 
|
56  | 
fun chk_prem rec_hd (Const("op &",_) $ _ $ _) =
 | 
|
| 1465 | 57  | 
error"Premises may not be conjuctive"  | 
| 923 | 58  | 
  | chk_prem rec_hd (Const("op :",_) $ t $ X) = 
 | 
| 1465 | 59  | 
deny (Logic.occs(rec_hd,t)) "Recursion term on left of member symbol"  | 
| 923 | 60  | 
| chk_prem rec_hd t =  | 
| 1465 | 61  | 
deny (Logic.occs(rec_hd,t)) "Recursion term in side formula";  | 
| 923 | 62  | 
|
63  | 
(*Return the conclusion of a rule, of the form t:X*)  | 
|
64  | 
fun rule_concl rl =  | 
|
65  | 
    let val Const("Trueprop",_) $ (Const("op :",_) $ t $ X) = 
 | 
|
| 1465 | 66  | 
Logic.strip_imp_concl rl  | 
| 923 | 67  | 
in (t,X) end;  | 
68  | 
||
69  | 
(*As above, but return error message if bad*)  | 
|
70  | 
fun rule_concl_msg sign rl = rule_concl rl  | 
|
71  | 
    handle Bind => error ("Ill-formed conclusion of introduction rule: " ^ 
 | 
|
| 1465 | 72  | 
Sign.string_of_term sign rl);  | 
| 923 | 73  | 
|
74  | 
(*For simplifying the elimination rule*)  | 
|
75  | 
val sumprod_free_SEs =  | 
|
76  | 
Pair_inject ::  | 
|
77  | 
map make_elim [(*Inl_neq_Inr, Inr_neq_Inl, Inl_inject, Inr_inject*)];  | 
|
78  | 
||
79  | 
(*For deriving cases rules.  | 
|
80  | 
read_instantiate replaces a propositional variable by a formula variable*)  | 
|
81  | 
val equals_CollectD =  | 
|
82  | 
    read_instantiate [("W","?Q")]
 | 
|
83  | 
(make_elim (equalityD1 RS subsetD RS CollectD));  | 
|
84  | 
||
85  | 
(*Delete needless equality assumptions*)  | 
|
86  | 
val refl_thin = prove_goal HOL.thy "!!P. [| a=a; P |] ==> P"  | 
|
87  | 
(fn _ => [assume_tac 1]);  | 
|
88  | 
||
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
89  | 
(*Includes rules for Suc and Pair since they are common constructions*)  | 
| 1433 | 90  | 
val elim_rls = [asm_rl, FalseE, (*Suc_neq_Zero, Zero_neq_Suc,  | 
| 1465 | 91  | 
make_elim Suc_inject, *)  | 
92  | 
refl_thin, conjE, exE, disjE];  | 
|
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
93  | 
|
| 923 | 94  | 
end;  |