| author | wenzelm | 
| Sat, 14 Jan 2017 20:39:16 +0100 | |
| changeset 64891 | d047004c1109 | 
| parent 63856 | 0db1481c1ec1 | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 54701 | 1 | (* Title: HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML | 
| 54397 | 2 | Author: Dmitriy Traytel, TU Muenchen | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 54397 | 4 | Copyright 2012, 2013 | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 5 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 6 | Library for wrapping existing freely generated type's constructors. | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 7 | *) | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 8 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 9 | signature CTR_SUGAR_UTIL = | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 10 | sig | 
| 57527 | 11 |   val map_prod: ('a -> 'b) -> ('c -> 'd) -> 'a * 'c -> 'b * 'd
 | 
| 12 | ||
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 13 | val seq_conds: (bool -> 'a -> 'b) -> int -> int -> 'a list -> 'b list | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 14 | val transpose: 'a list list -> 'a list list | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 15 | val pad_list: 'a -> int -> 'a list -> 'a list | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 16 | val splice: 'a list -> 'a list -> 'a list | 
| 55480 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 17 |   val permute_like_unique: ('a * 'b -> bool) -> 'a list -> 'b list -> 'c list -> 'c list
 | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 18 |   val permute_like: ('a * 'a -> bool) -> 'a list -> 'a list -> 'b list -> 'b list
 | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 19 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 20 | val mk_names: int -> string -> string list | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 21 | val mk_fresh_names: Proof.context -> int -> string -> string list * Proof.context | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 22 | val mk_TFrees': sort list -> Proof.context -> typ list * Proof.context | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 23 | val mk_TFrees: int -> Proof.context -> typ list * Proof.context | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 24 | val mk_Frees': string -> typ list -> Proof.context -> | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 25 | (term list * (string * typ) list) * Proof.context | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 26 | val mk_Freess': string -> typ list list -> Proof.context -> | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 27 | (term list list * (string * typ) list list) * Proof.context | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 28 | val mk_Frees: string -> typ list -> Proof.context -> term list * Proof.context | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 29 | val mk_Freess: string -> typ list list -> Proof.context -> term list list * Proof.context | 
| 58227 | 30 | val dest_TFree_or_TVar: typ -> string * sort | 
| 58234 | 31 | val resort_tfree_or_tvar: sort -> typ -> typ | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 32 | val variant_types: string list -> sort list -> Proof.context -> | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 33 | (string * sort) list * Proof.context | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 34 | val variant_tfrees: string list -> Proof.context -> typ list * Proof.context | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 35 | |
| 58284 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 36 | val base_name_of_typ: typ -> string | 
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 37 | val name_of_const: string -> (typ -> typ) -> term -> string | 
| 57700 | 38 | |
| 54435 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 39 | val typ_subst_nonatomic: (typ * typ) list -> typ -> typ | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 40 | val subst_nonatomic_types: (typ * typ) list -> term -> term | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 41 | |
| 55535 
10194808430d
name derivations in 'primrec' for code extraction from proof terms
 blanchet parents: 
55480diff
changeset | 42 | val lhs_head_of : thm -> term | 
| 55471 | 43 | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 44 | val mk_predT: typ list -> typ | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 45 | val mk_pred1T: typ -> typ | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 46 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 47 | val mk_disjIN: int -> int -> thm | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 48 | |
| 63856 | 49 | val mk_abs_def: thm -> thm | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 50 | val mk_unabs_def: int -> thm -> thm | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 51 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 52 | val mk_IfN: typ -> term list -> term list -> term | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 53 | val mk_Trueprop_eq: term * term -> term | 
| 57567 | 54 | val mk_Trueprop_mem: term * term -> term | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 55 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 56 | val rapp: term -> term -> term | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 57 | |
| 60728 | 58 | val rtac: Proof.context -> thm -> int -> tactic | 
| 59 | val etac: Proof.context -> thm -> int -> tactic | |
| 60 | val dtac: Proof.context -> thm -> int -> tactic | |
| 61 | ||
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 62 | val list_all_free: term list -> term -> term | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 63 | val list_exists_free: term list -> term -> term | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 64 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 65 | val fo_match: Proof.context -> term -> term -> Type.tyenv * Envir.tenv | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 66 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 67 | val unfold_thms: Proof.context -> thm list -> thm -> thm | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 68 | |
| 57633 | 69 | val name_noted_thms: string -> string -> (string * thm list) list -> (string * thm list) list | 
| 57629 
e88b5f59cade
use the noted theorem whenever possible, because it has a named derivation (leading to cleaner proof terms)
 blanchet parents: 
57567diff
changeset | 70 | val substitute_noted_thm: (string * thm list) list -> morphism | 
| 
e88b5f59cade
use the noted theorem whenever possible, because it has a named derivation (leading to cleaner proof terms)
 blanchet parents: 
57567diff
changeset | 71 | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 72 | val standard_binding: binding | 
| 55469 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 blanchet parents: 
54701diff
changeset | 73 | val parse_binding_colon: binding parser | 
| 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 blanchet parents: 
54701diff
changeset | 74 | val parse_opt_binding_colon: binding parser | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 75 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 76 | val ss_only: thm list -> Proof.context -> Proof.context | 
| 54540 | 77 | |
| 59271 | 78 | (*parameterized thms*) | 
| 79 | val eqTrueI: thm | |
| 80 | val eqFalseI: thm | |
| 81 | ||
| 54540 | 82 |   val WRAP: ('a -> tactic) -> ('a -> tactic) -> 'a list -> tactic -> tactic
 | 
| 83 |   val WRAP': ('a -> int -> tactic) -> ('a -> int -> tactic) -> 'a list -> (int -> tactic) -> int ->
 | |
| 84 | tactic | |
| 85 |   val CONJ_WRAP_GEN: tactic -> ('a -> tactic) -> 'a list -> tactic
 | |
| 86 |   val CONJ_WRAP_GEN': (int -> tactic) -> ('a -> int -> tactic) -> 'a list -> int -> tactic
 | |
| 87 |   val CONJ_WRAP: ('a -> tactic) -> 'a list -> tactic
 | |
| 88 |   val CONJ_WRAP': ('a -> int -> tactic) -> 'a list -> int -> tactic
 | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 89 | end; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 90 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 91 | structure Ctr_Sugar_Util : CTR_SUGAR_UTIL = | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 92 | struct | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 93 | |
| 58220 | 94 | fun map_prod f g (x, y) = (f x, g y); | 
| 57527 | 95 | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 96 | fun seq_conds f n k xs = | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 97 | if k = n then | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 98 | map (f false) (take (k - 1) xs) | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 99 | else | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 100 | let val (negs, pos) = split_last (take k xs) in | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 101 | map (f false) negs @ [f true pos] | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 102 | end; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 103 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 104 | fun transpose [] = [] | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 105 | | transpose ([] :: xss) = transpose xss | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 106 | | transpose xss = map hd xss :: transpose (map tl xss); | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 107 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 108 | fun pad_list x n xs = xs @ replicate (n - length xs) x; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 109 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 110 | fun splice xs ys = flat (map2 (fn x => fn y => [x, y]) xs ys); | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 111 | |
| 55480 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 112 | fun permute_like_unique eq xs xs' ys = | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 113 | map (nth ys o (fn y => find_index (fn x => eq (x, y)) xs)) xs'; | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 114 | |
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 115 | fun fresh eq x names = | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 116 | (case AList.lookup eq names x of | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 117 | NONE => ((x, 0), (x, 0) :: names) | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 118 | | SOME n => ((x, n + 1), AList.update eq (x, n + 1) names)); | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 119 | |
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 120 | fun deambiguate eq xs = fst (fold_map (fresh eq) xs []); | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 121 | |
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 122 | fun permute_like eq xs xs' = | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55471diff
changeset | 123 | permute_like_unique (eq_pair eq (op =)) (deambiguate eq xs) (deambiguate eq xs'); | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 124 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 125 | fun mk_names n x = if n = 1 then [x] else map (fn i => x ^ string_of_int i) (1 upto n); | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 126 | fun mk_fresh_names ctxt = (fn xs => Variable.variant_fixes xs ctxt) oo mk_names; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 127 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 128 | val mk_TFrees' = apfst (map TFree) oo Variable.invent_types; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 129 | |
| 56254 | 130 | fun mk_TFrees n = mk_TFrees' (replicate n @{sort type});
 | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 131 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 132 | fun mk_Frees' x Ts ctxt = mk_fresh_names ctxt (length Ts) x |>> (fn xs => `(map Free) (xs ~~ Ts)); | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58284diff
changeset | 133 | fun mk_Freess' x Tss = @{fold_map 2} mk_Frees' (mk_names (length Tss) x) Tss #>> split_list;
 | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 134 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 135 | fun mk_Frees x Ts ctxt = mk_fresh_names ctxt (length Ts) x |>> (fn xs => map2 (curry Free) xs Ts); | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58284diff
changeset | 136 | fun mk_Freess x Tss = @{fold_map 2} mk_Frees (mk_names (length Tss) x) Tss;
 | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 137 | |
| 58227 | 138 | fun dest_TFree_or_TVar (TFree sS) = sS | 
| 139 | | dest_TFree_or_TVar (TVar ((s, _), S)) = (s, S) | |
| 140 | | dest_TFree_or_TVar _ = error "Invalid type argument"; | |
| 141 | ||
| 58234 | 142 | fun resort_tfree_or_tvar S (TFree (s, _)) = TFree (s, S) | 
| 143 | | resort_tfree_or_tvar S (TVar (x, _)) = TVar (x, S); | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 144 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 145 | fun ensure_prefix pre s = s |> not (String.isPrefix pre s) ? prefix pre; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 146 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 147 | fun variant_types ss Ss ctxt = | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 148 | let | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 149 | val (tfrees, _) = | 
| 59650 | 150 |       @{fold_map 2} (fn s => fn S => Name.variant s #> apfst (rpair S))
 | 
| 151 | ss Ss (Variable.names_of ctxt); | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 152 | val ctxt' = fold (Variable.declare_constraints o Logic.mk_type o TFree) tfrees ctxt; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 153 | in (tfrees, ctxt') end; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 154 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 155 | fun variant_tfrees ss = | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 156 | apfst (map TFree) o | 
| 56254 | 157 |     variant_types (map (ensure_prefix "'") ss) (replicate (length ss) @{sort type});
 | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 158 | |
| 58284 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 159 | fun add_components_of_typ (Type (s, Ts)) = | 
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 160 | cons (Long_Name.base_name s) #> fold_rev add_components_of_typ Ts | 
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 161 | | add_components_of_typ _ = I; | 
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 162 | |
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 163 | fun base_name_of_typ T = space_implode "_" (add_components_of_typ T []); | 
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 164 | |
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 165 | fun suffix_with_type s (Type (_, Ts)) = | 
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 166 | space_implode "_" (s :: fold_rev add_components_of_typ Ts []) | 
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 167 | | suffix_with_type s _ = s; | 
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 168 | |
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 169 | fun name_of_const what get_fcT t = | 
| 57700 | 170 | (case head_of t of | 
| 58284 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 171 | Const (s, T) => suffix_with_type s (get_fcT T) | 
| 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 blanchet parents: 
58234diff
changeset | 172 | | Free (s, T) => suffix_with_type s (get_fcT T) | 
| 57700 | 173 |   | _ => error ("Cannot extract name of " ^ what));
 | 
| 174 | ||
| 54435 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 175 | (*Replace each Ti by Ui (starting from the leaves); inst = [(T1, U1), ..., (Tn, Un)].*) | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 176 | fun typ_subst_nonatomic [] = I | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 177 | | typ_subst_nonatomic inst = | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 178 | let | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 179 | fun subst (Type (s, Ts)) = perhaps (AList.lookup (op =) inst) (Type (s, map subst Ts)) | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 180 | | subst T = perhaps (AList.lookup (op =) inst) T; | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 181 | in subst end; | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 182 | |
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 183 | fun subst_nonatomic_types [] = I | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 184 | | subst_nonatomic_types inst = map_types (typ_subst_nonatomic inst); | 
| 
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
 blanchet parents: 
54397diff
changeset | 185 | |
| 59582 | 186 | fun lhs_head_of thm = | 
| 187 | Term.head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm)))); | |
| 55471 | 188 | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 189 | fun mk_predT Ts = Ts ---> HOLogic.boolT; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 190 | fun mk_pred1T T = mk_predT [T]; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 191 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 192 | fun mk_disjIN 1 1 = @{thm TrueE[OF TrueI]}
 | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 193 | | mk_disjIN _ 1 = disjI1 | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 194 | | mk_disjIN 2 2 = disjI2 | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 195 | | mk_disjIN n m = (mk_disjIN (n - 1) (m - 1)) RS disjI2; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 196 | |
| 63856 | 197 | fun mk_abs_def thm = Drule.abs_def (thm RS eq_reflection handle THM _ => thm); | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 198 | fun mk_unabs_def n = funpow n (fn thm => thm RS fun_cong); | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 199 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 200 | fun mk_IfN _ _ [t] = t | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 201 | | mk_IfN T (c :: cs) (t :: ts) = | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 202 |     Const (@{const_name If}, HOLogic.boolT --> T --> T --> T) $ c $ t $ mk_IfN T cs ts;
 | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 203 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 204 | val mk_Trueprop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq; | 
| 57567 | 205 | val mk_Trueprop_mem = HOLogic.mk_Trueprop o HOLogic.mk_mem; | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 206 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 207 | fun rapp u t = betapply (t, u); | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 208 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 209 | fun list_quant_free quant_const = | 
| 54491 | 210 | fold_rev (fn Free (xT as (_, T)) => fn P => quant_const T $ Term.absfree xT P); | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 211 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 212 | val list_all_free = list_quant_free HOLogic.all_const; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 213 | val list_exists_free = list_quant_free HOLogic.exists_const; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 214 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 215 | fun fo_match ctxt t pat = | 
| 60788 | 216 | let val thy = Proof_Context.theory_of ctxt | 
| 217 | in Pattern.first_order_match thy (pat, t) (Vartab.empty, Vartab.empty) end; | |
| 54540 | 218 | |
| 63223 | 219 | val unfold_thms = Local_Defs.unfold0; | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 220 | |
| 57633 | 221 | fun name_noted_thms _ _ [] = [] | 
| 222 | | name_noted_thms qualifier base ((local_name, thms) :: noted) = | |
| 223 | if Long_Name.base_name local_name = base then | |
| 224 | (local_name, | |
| 59877 | 225 | map_index (uncurry (fn j => Thm.name_derivation (Long_Name.append qualifier base ^ | 
| 57633 | 226 | (if can the_single thms then "" else "_" ^ string_of_int (j + 1))))) thms) :: noted | 
| 227 | else | |
| 228 | ((local_name, thms) :: name_noted_thms qualifier base noted); | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 229 | |
| 57629 
e88b5f59cade
use the noted theorem whenever possible, because it has a named derivation (leading to cleaner proof terms)
 blanchet parents: 
57567diff
changeset | 230 | fun substitute_noted_thm noted = | 
| 57630 
04ab38720b50
use termtab instead of (perhaps overly sensitive) thmtab
 blanchet parents: 
57629diff
changeset | 231 | let val tab = fold (fold (Termtab.default o `Thm.full_prop_of) o snd) noted Termtab.empty in | 
| 
04ab38720b50
use termtab instead of (perhaps overly sensitive) thmtab
 blanchet parents: 
57629diff
changeset | 232 | Morphism.thm_morphism "Ctr_Sugar_Util.substitute_noted_thm" | 
| 57631 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 blanchet parents: 
57630diff
changeset | 233 | (perhaps (Termtab.lookup tab o Thm.full_prop_of) o Drule.zero_var_indexes) | 
| 57633 | 234 | end; | 
| 57629 
e88b5f59cade
use the noted theorem whenever possible, because it has a named derivation (leading to cleaner proof terms)
 blanchet parents: 
57567diff
changeset | 235 | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 236 | (* The standard binding stands for a name generated following the canonical convention (e.g., | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 237 | "is_Nil" from "Nil"). In contrast, the empty binding is either the standard binding or no | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 238 | binding at all, depending on the context. *) | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 239 | val standard_binding = @{binding _};
 | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 240 | |
| 57091 | 241 | val parse_binding_colon = Parse.binding --| @{keyword ":"};
 | 
| 55469 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 blanchet parents: 
54701diff
changeset | 242 | val parse_opt_binding_colon = Scan.optional parse_binding_colon Binding.empty; | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 243 | |
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 244 | fun ss_only thms ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps thms; | 
| 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 245 | |
| 59271 | 246 | val eqTrueI = @{thm iffD2[OF eq_True]};
 | 
| 247 | val eqFalseI =  @{thm iffD2[OF eq_False]};
 | |
| 248 | ||
| 54540 | 249 | (*Tactical WRAP surrounds a static given tactic (core) with two deterministic chains of tactics*) | 
| 250 | fun WRAP gen_before gen_after xs core_tac = | |
| 251 | fold_rev (fn x => fn tac => gen_before x THEN tac THEN gen_after x) xs core_tac; | |
| 252 | ||
| 253 | fun WRAP' gen_before gen_after xs core_tac = | |
| 254 | fold_rev (fn x => fn tac => gen_before x THEN' tac THEN' gen_after x) xs core_tac; | |
| 255 | ||
| 256 | fun CONJ_WRAP_GEN conj_tac gen_tac xs = | |
| 257 | let val (butlast, last) = split_last xs; | |
| 258 | in WRAP (fn thm => conj_tac THEN gen_tac thm) (K all_tac) butlast (gen_tac last) end; | |
| 259 | ||
| 260 | fun CONJ_WRAP_GEN' conj_tac gen_tac xs = | |
| 261 | let val (butlast, last) = split_last xs; | |
| 262 | in WRAP' (fn thm => conj_tac THEN' gen_tac thm) (K (K all_tac)) butlast (gen_tac last) end; | |
| 263 | ||
| 60752 | 264 | fun CONJ_WRAP gen_tac = CONJ_WRAP_GEN (resolve0_tac [conjI] 1) gen_tac; | 
| 265 | fun CONJ_WRAP' gen_tac = CONJ_WRAP_GEN' (resolve0_tac [conjI]) gen_tac; | |
| 54540 | 266 | |
| 60728 | 267 | fun rtac ctxt thm = resolve_tac ctxt [thm]; | 
| 268 | fun etac ctxt thm = eresolve_tac ctxt [thm]; | |
| 269 | fun dtac ctxt thm = dresolve_tac ctxt [thm]; | |
| 270 | ||
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: diff
changeset | 271 | end; |