| author | hoelzl | 
| Mon, 19 May 2014 13:44:13 +0200 | |
| changeset 56994 | 8d5e5ec1cac3 | 
| parent 56245 | 84fc7dfa3cd4 | 
| child 58634 | 9f10d82e8188 | 
| permissions | -rw-r--r-- | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1 | (* Title: HOL/Tools/Nitpick/nitpick_preproc.ML | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 2 | Author: Jasmin Blanchette, TU Muenchen | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 3 | Copyright 2008, 2009, 2010 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 4 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 5 | Nitpick's HOL preprocessor. | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 6 | *) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 7 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 8 | signature NITPICK_PREPROC = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 9 | sig | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 10 | type hol_context = Nitpick_HOL.hol_context | 
| 38169 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 11 | val preprocess_formulas : | 
| 41052 
3db267a01c1d
remove the "fin_fun" optimization in Nitpick -- it was always a hack and didn't help much
 blanchet parents: 
41001diff
changeset | 12 | hol_context -> term list -> term | 
| 41803 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41793diff
changeset | 13 | -> term list * term list * term list * bool * bool * bool | 
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35807diff
changeset | 14 | end; | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 15 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 16 | structure Nitpick_Preproc : NITPICK_PREPROC = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 17 | struct | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 18 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 19 | open Nitpick_Util | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 20 | open Nitpick_HOL | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35408diff
changeset | 21 | open Nitpick_Mono | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 22 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 23 | fun is_positive_existential polar quant_s = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 24 |   (polar = Pos andalso quant_s = @{const_name Ex}) orelse
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 25 |   (polar = Neg andalso quant_s <> @{const_name Ex})
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 26 | |
| 37271 | 27 | val is_descr = | 
| 39360 
cdf2c3341422
eliminate more clutter related to "fast_descrs" optimization
 blanchet parents: 
39359diff
changeset | 28 |   member (op =) [@{const_name The}, @{const_name Eps}, @{const_name safe_The}]
 | 
| 37271 | 29 | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 30 | (** Binary coding of integers **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 31 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 32 | (* If a formula contains a numeral whose absolute value is more than this | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 33 | threshold, the unary coding is likely not to work well and we prefer the | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 34 | binary coding. *) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 35 | val binary_int_threshold = 3 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 36 | |
| 35718 | 37 | val may_use_binary_ints = | 
| 38 | let | |
| 56245 | 39 |     fun aux def (Const (@{const_name Pure.eq}, _) $ t1 $ t2) =
 | 
| 35718 | 40 | aux def t1 andalso aux false t2 | 
| 56245 | 41 |       | aux def (@{const Pure.imp} $ t1 $ t2) = aux false t1 andalso aux def t2
 | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 42 |       | aux def (Const (@{const_name HOL.eq}, _) $ t1 $ t2) =
 | 
| 35718 | 43 | aux def t1 andalso aux false t2 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38240diff
changeset | 44 |       | aux def (@{const HOL.implies} $ t1 $ t2) = aux false t1 andalso aux def t2
 | 
| 35718 | 45 | | aux def (t1 $ t2) = aux def t1 andalso aux def t2 | 
| 46 | | aux def (t as Const (s, _)) = | |
| 47 |         (not def orelse t <> @{const Suc}) andalso
 | |
| 47909 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 blanchet parents: 
47754diff
changeset | 48 | not (member (op =) | 
| 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 blanchet parents: 
47754diff
changeset | 49 |                [@{const_name Nitpick.Abs_Frac}, @{const_name Nitpick.Rep_Frac},
 | 
| 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 blanchet parents: 
47754diff
changeset | 50 |                 @{const_name Nitpick.nat_gcd}, @{const_name Nitpick.nat_lcm},
 | 
| 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 blanchet parents: 
47754diff
changeset | 51 |                 @{const_name Nitpick.Frac}, @{const_name Nitpick.norm_frac}] s)
 | 
| 35718 | 52 | | aux def (Abs (_, _, t')) = aux def t' | 
| 53 | | aux _ _ = true | |
| 54 | in aux end | |
| 55 | val should_use_binary_ints = | |
| 56 | let | |
| 57 | fun aux (t1 $ t2) = aux t1 orelse aux t2 | |
| 58 | | aux (Const (s, T)) = | |
| 59 |         ((s = @{const_name times} orelse s = @{const_name div}) andalso
 | |
| 60 | is_integer_type (body_type T)) orelse | |
| 61 | (String.isPrefix numeral_prefix s andalso | |
| 62 | let val n = the (Int.fromString (unprefix numeral_prefix s)) in | |
| 63 | n < ~ binary_int_threshold orelse n > binary_int_threshold | |
| 64 | end) | |
| 65 | | aux (Abs (_, _, t')) = aux t' | |
| 66 | | aux _ = false | |
| 67 | in aux end | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 68 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 69 | (** Uncurrying **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 70 | |
| 37256 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 71 | fun add_to_uncurry_table ctxt t = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 72 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 73 | fun aux (t1 $ t2) args table = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 74 | let val table = aux t2 [] table in aux t1 (t2 :: args) table end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 75 | | aux (Abs (_, _, t')) _ table = aux t' [] table | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 76 | | aux (t as Const (x as (s, _))) args table = | 
| 55889 | 77 | if is_built_in_const x orelse is_nonfree_constr ctxt x orelse | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 78 |            is_sel s orelse s = @{const_name Sigma} then
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 79 | table | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 80 | else | 
| 36384 | 81 | Termtab.map_default (t, 65536) (Integer.min (length args)) table | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 82 | | aux _ _ table = table | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 83 | in aux t [] end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 84 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 85 | fun uncurry_prefix_for k j = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 86 | uncurry_prefix ^ string_of_int k ^ "@" ^ string_of_int j ^ name_sep | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 87 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 88 | fun uncurry_term table t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 89 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 90 | fun aux (t1 $ t2) args = aux t1 (aux t2 [] :: args) | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 91 | | aux (Abs (s, T, t')) args = s_betapplys [] (Abs (s, T, aux t' []), args) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 92 | | aux (t as Const (s, T)) args = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 93 | (case Termtab.lookup table t of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 94 | SOME n => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 95 | if n >= 2 then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 96 | let | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 97 | val arg_Ts = strip_n_binders n T |> fst | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 98 | val j = | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 99 | if is_iterator_type (hd arg_Ts) then | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 100 | 1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 101 | else case find_index (not_equal bool_T) arg_Ts of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 102 | ~1 => n | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 103 | | j => j | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 104 | val ((before_args, tuple_args), after_args) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 105 | args |> chop n |>> chop j | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 106 | val ((before_arg_Ts, tuple_arg_Ts), rest_T) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 107 | T |> strip_n_binders n |>> chop j | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 108 | val tuple_T = HOLogic.mk_tupleT tuple_arg_Ts | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 109 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 110 | if n - j < 2 then | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 111 | s_betapplys [] (t, args) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 112 | else | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 113 | s_betapplys [] | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 114 | (Const (uncurry_prefix_for (n - j) j ^ s, | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 115 | before_arg_Ts ---> tuple_T --> rest_T), | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 116 | before_args @ [mk_flat_tuple tuple_T tuple_args] @ | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 117 | after_args) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 118 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 119 | else | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 120 | s_betapplys [] (t, args) | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 121 | | NONE => s_betapplys [] (t, args)) | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 122 | | aux t args = s_betapplys [] (t, args) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 123 | in aux t [] end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 124 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 125 | (** Boxing **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 126 | |
| 55888 | 127 | fun box_fun_and_pair_in_term (hol_ctxt as {ctxt, ...}) def orig_t =
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 128 | let | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35408diff
changeset | 129 |     fun box_relational_operator_type (Type (@{type_name fun}, Ts)) =
 | 
| 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35408diff
changeset | 130 |         Type (@{type_name fun}, map box_relational_operator_type Ts)
 | 
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38170diff
changeset | 131 |       | box_relational_operator_type (Type (@{type_name prod}, Ts)) =
 | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38170diff
changeset | 132 |         Type (@{type_name prod}, map (box_type hol_ctxt InPair) Ts)
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 133 | | box_relational_operator_type T = T | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 134 | fun add_boxed_types_for_var (z as (_, T)) (T', t') = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 135 | case t' of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 136 | Var z' => z' = z ? insert (op =) T' | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 137 |       | Const (@{const_name Pair}, _) $ t1 $ t2 =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 138 | (case T' of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 139 | Type (_, [T1, T2]) => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 140 | fold (add_boxed_types_for_var z) [(T1, t1), (T2, t2)] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 141 |          | _ => raise TYPE ("Nitpick_Preproc.box_fun_and_pair_in_term.\
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 142 | \add_boxed_types_for_var", [T'], [])) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 143 | | _ => exists_subterm (curry (op =) (Var z)) t' ? insert (op =) T | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 144 | fun box_var_in_def new_Ts old_Ts t (z as (_, T)) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 145 | case t of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 146 |         @{const Trueprop} $ t1 => box_var_in_def new_Ts old_Ts t1 z
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 147 | | Const (s0, _) $ t1 $ _ => | 
| 56245 | 148 |         if s0 = @{const_name Pure.eq} orelse s0 = @{const_name HOL.eq} then
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 149 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 150 | val (t', args) = strip_comb t1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 151 | val T' = fastype_of1 (new_Ts, do_term new_Ts old_Ts Neut t') | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 152 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 153 | case fold (add_boxed_types_for_var z) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 154 | (fst (strip_n_binders (length args) T') ~~ args) [] of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 155 | [T''] => T'' | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 156 | | _ => T | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 157 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 158 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 159 | T | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 160 | | _ => T | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 161 | and do_quantifier new_Ts old_Ts polar quant_s quant_T abs_s abs_T t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 162 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 163 | val abs_T' = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 164 | if polar = Neut orelse is_positive_existential polar quant_s then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 165 | box_type hol_ctxt InFunLHS abs_T | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 166 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 167 | abs_T | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 168 | val body_T = body_type quant_T | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 169 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 170 | Const (quant_s, (abs_T' --> body_T) --> body_T) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 171 | $ Abs (abs_s, abs_T', | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 172 | t |> do_term (abs_T' :: new_Ts) (abs_T :: old_Ts) polar) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 173 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 174 | and do_equals new_Ts old_Ts s0 T0 t1 t2 = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 175 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 176 | val (t1, t2) = pairself (do_term new_Ts old_Ts Neut) (t1, t2) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 177 | val (T1, T2) = pairself (curry fastype_of1 new_Ts) (t1, t2) | 
| 37704 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 blanchet parents: 
37678diff
changeset | 178 | val T = if def then T1 | 
| 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 blanchet parents: 
37678diff
changeset | 179 | else [T1, T2] |> sort (int_ord o pairself size_of_typ) |> hd | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 180 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 181 | list_comb (Const (s0, T --> T --> body_type T0), | 
| 35384 | 182 | map2 (coerce_term hol_ctxt new_Ts T) [T1, T2] [t1, t2]) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 183 | end | 
| 37271 | 184 | and do_descr s T = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 185 | let val T1 = box_type hol_ctxt InFunLHS (range_type T) in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 186 | Const (s, (T1 --> bool_T) --> T1) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 187 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 188 | and do_term new_Ts old_Ts polar t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 189 | case t of | 
| 56245 | 190 |         Const (s0 as @{const_name Pure.all}, T0) $ Abs (s1, T1, t1) =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 191 | do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1 | 
| 56245 | 192 |       | Const (s0 as @{const_name Pure.eq}, T0) $ t1 $ t2 =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 193 | do_equals new_Ts old_Ts s0 T0 t1 t2 | 
| 56245 | 194 |       | @{const Pure.imp} $ t1 $ t2 =>
 | 
| 195 |         @{const Pure.imp} $ do_term new_Ts old_Ts (flip_polarity polar) t1
 | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 196 | $ do_term new_Ts old_Ts polar t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 197 |       | @{const Pure.conjunction} $ t1 $ t2 =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 198 |         @{const Pure.conjunction} $ do_term new_Ts old_Ts polar t1
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 199 | $ do_term new_Ts old_Ts polar t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 200 |       | @{const Trueprop} $ t1 =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 201 |         @{const Trueprop} $ do_term new_Ts old_Ts polar t1
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 202 |       | @{const Not} $ t1 =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 203 |         @{const Not} $ do_term new_Ts old_Ts (flip_polarity polar) t1
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 204 |       | Const (s0 as @{const_name All}, T0) $ Abs (s1, T1, t1) =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 205 | do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 206 |       | Const (s0 as @{const_name Ex}, T0) $ Abs (s1, T1, t1) =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 207 | do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1 | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 208 |       | Const (s0 as @{const_name HOL.eq}, T0) $ t1 $ t2 =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 209 | do_equals new_Ts old_Ts s0 T0 t1 t2 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 210 |       | @{const HOL.conj} $ t1 $ t2 =>
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 211 |         @{const HOL.conj} $ do_term new_Ts old_Ts polar t1
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 212 | $ do_term new_Ts old_Ts polar t2 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 213 |       | @{const HOL.disj} $ t1 $ t2 =>
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 214 |         @{const HOL.disj} $ do_term new_Ts old_Ts polar t1
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 215 | $ do_term new_Ts old_Ts polar t2 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38240diff
changeset | 216 |       | @{const HOL.implies} $ t1 $ t2 =>
 | 
| 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38240diff
changeset | 217 |         @{const HOL.implies} $ do_term new_Ts old_Ts (flip_polarity polar) t1
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 218 | $ do_term new_Ts old_Ts polar t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 219 | | Const (x as (s, T)) => | 
| 37271 | 220 | if is_descr s then | 
| 221 | do_descr s T | |
| 222 | else | |
| 223 |           Const (s, if s = @{const_name converse} orelse
 | |
| 224 |                        s = @{const_name trancl} then
 | |
| 225 | box_relational_operator_type T | |
| 226 | else if String.isPrefix quot_normal_prefix s then | |
| 227 | let val T' = box_type hol_ctxt InFunLHS (domain_type T) in | |
| 228 | T' --> T' | |
| 229 | end | |
| 55888 | 230 | else if is_built_in_const x orelse | 
| 37271 | 231 |                             s = @{const_name Sigma} then
 | 
| 232 | T | |
| 55889 | 233 | else if is_nonfree_constr ctxt x then | 
| 37271 | 234 | box_type hol_ctxt InConstr T | 
| 235 | else if is_sel s orelse is_rep_fun ctxt x then | |
| 236 | box_type hol_ctxt InSel T | |
| 237 | else | |
| 238 | box_type hol_ctxt InExpr T) | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 239 | | t1 $ Abs (s, T, t2') => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 240 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 241 | val t1 = do_term new_Ts old_Ts Neut t1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 242 | val T1 = fastype_of1 (new_Ts, t1) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 243 | val (s1, Ts1) = dest_Type T1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 244 | val T' = hd (snd (dest_Type (hd Ts1))) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 245 | val t2 = Abs (s, T', do_term (T' :: new_Ts) (T :: old_Ts) Neut t2') | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 246 | val T2 = fastype_of1 (new_Ts, t2) | 
| 35384 | 247 | val t2 = coerce_term hol_ctxt new_Ts (hd Ts1) T2 t2 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 248 | in | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 249 |           s_betapply new_Ts (if s1 = @{type_name fun} then
 | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 250 | t1 | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 251 | else | 
| 55888 | 252 | select_nth_constr_arg ctxt | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 253 |                                    (@{const_name FunBox},
 | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 254 |                                     Type (@{type_name fun}, Ts1) --> T1) t1 0
 | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 255 |                                    (Type (@{type_name fun}, Ts1)), t2)
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 256 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 257 | | t1 $ t2 => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 258 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 259 | val t1 = do_term new_Ts old_Ts Neut t1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 260 | val T1 = fastype_of1 (new_Ts, t1) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 261 | val (s1, Ts1) = dest_Type T1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 262 | val t2 = do_term new_Ts old_Ts Neut t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 263 | val T2 = fastype_of1 (new_Ts, t2) | 
| 35384 | 264 | val t2 = coerce_term hol_ctxt new_Ts (hd Ts1) T2 t2 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 265 | in | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 266 |           s_betapply new_Ts (if s1 = @{type_name fun} then
 | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 267 | t1 | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 268 | else | 
| 55888 | 269 | select_nth_constr_arg ctxt | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 270 |                                    (@{const_name FunBox},
 | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 271 |                                     Type (@{type_name fun}, Ts1) --> T1) t1 0
 | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 272 |                                    (Type (@{type_name fun}, Ts1)), t2)
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 273 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 274 | | Free (s, T) => Free (s, box_type hol_ctxt InExpr T) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 275 | | Var (z as (x, T)) => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 276 | Var (x, if def then box_var_in_def new_Ts old_Ts orig_t z | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 277 | else box_type hol_ctxt InExpr T) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 278 | | Bound _ => t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 279 | | Abs (s, T, t') => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 280 | Abs (s, T, do_term (T :: new_Ts) (T :: old_Ts) Neut t') | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 281 | in do_term [] [] Pos orig_t end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 282 | |
| 46102 | 283 | (** Destruction of set membership and comprehensions **) | 
| 284 | ||
| 285 | fun destroy_set_Collect (Const (@{const_name Set.member}, _) $ t1
 | |
| 286 |                          $ (Const (@{const_name Collect}, _) $ t2)) =
 | |
| 287 | destroy_set_Collect (t2 $ t1) | |
| 288 | | destroy_set_Collect (t1 $ t2) = | |
| 289 | destroy_set_Collect t1 $ destroy_set_Collect t2 | |
| 290 | | destroy_set_Collect (Abs (s, T, t')) = Abs (s, T, destroy_set_Collect t') | |
| 291 | | destroy_set_Collect t = t | |
| 292 | ||
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 293 | (** Destruction of constructors **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 294 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 295 | val val_var_prefix = nitpick_prefix ^ "v" | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 296 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 297 | fun fresh_value_var Ts k n j t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 298 | Var ((val_var_prefix ^ nat_subscript (n - j), k), fastype_of1 (Ts, t)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 299 | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 300 | fun has_heavy_bounds_or_vars Ts t = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 301 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 302 | fun aux [] = false | 
| 46115 | 303 | | aux [T] = is_fun_or_set_type T orelse is_pair_type T | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 304 | | aux _ = true | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 305 | in aux (map snd (Term.add_vars t []) @ map (nth Ts) (loose_bnos t)) end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 306 | |
| 55888 | 307 | fun pull_out_constr_comb ({ctxt, ...} : hol_context) Ts relax k level t args
 | 
| 308 | seen = | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 309 | let val t_comb = list_comb (t, args) in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 310 | case t of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 311 | Const x => | 
| 55888 | 312 | if not relax andalso is_constr ctxt x andalso | 
| 46115 | 313 | not (is_fun_or_set_type (fastype_of1 (Ts, t_comb))) andalso | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 314 | has_heavy_bounds_or_vars Ts t_comb andalso | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 315 | not (loose_bvar (t_comb, level)) then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 316 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 317 | val (j, seen) = case find_index (curry (op =) t_comb) seen of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 318 | ~1 => (0, t_comb :: seen) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 319 | | j => (j, seen) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 320 | in (fresh_value_var Ts k (length seen) j t_comb, seen) end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 321 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 322 | (t_comb, seen) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 323 | | _ => (t_comb, seen) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 324 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 325 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 326 | fun equations_for_pulled_out_constrs mk_eq Ts k seen = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 327 | let val n = length seen in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 328 | map2 (fn j => fn t => mk_eq (fresh_value_var Ts k n j t, t)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 329 | (index_seq 0 n) seen | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 330 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 331 | |
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 332 | fun pull_out_universal_constrs hol_ctxt def t = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 333 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 334 | val k = maxidx_of_term t + 1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 335 | fun do_term Ts def t args seen = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 336 | case t of | 
| 56245 | 337 |         (t0 as Const (@{const_name Pure.eq}, _)) $ t1 $ t2 =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 338 | do_eq_or_imp Ts true def t0 t1 t2 seen | 
| 56245 | 339 |       | (t0 as @{const Pure.imp}) $ t1 $ t2 =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 340 | if def then (t, []) else do_eq_or_imp Ts false def t0 t1 t2 seen | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 341 |       | (t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 342 | do_eq_or_imp Ts true def t0 t1 t2 seen | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38240diff
changeset | 343 |       | (t0 as @{const HOL.implies}) $ t1 $ t2 =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 344 | do_eq_or_imp Ts false def t0 t1 t2 seen | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 345 | | Abs (s, T, t') => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 346 | let val (t', seen) = do_term (T :: Ts) def t' [] seen in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 347 | (list_comb (Abs (s, T, t'), args), seen) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 348 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 349 | | t1 $ t2 => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 350 | let val (t2, seen) = do_term Ts def t2 [] seen in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 351 | do_term Ts def t1 (t2 :: args) seen | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 352 | end | 
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 353 | | _ => pull_out_constr_comb hol_ctxt Ts def k 0 t args seen | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 354 | and do_eq_or_imp Ts eq def t0 t1 t2 seen = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 355 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 356 | val (t2, seen) = if eq andalso def then (t2, seen) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 357 | else do_term Ts false t2 [] seen | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 358 | val (t1, seen) = do_term Ts false t1 [] seen | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 359 | in (t0 $ t1 $ t2, seen) end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 360 | val (concl, seen) = do_term [] def t [] [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 361 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 362 | Logic.list_implies (equations_for_pulled_out_constrs Logic.mk_equals [] k | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 363 | seen, concl) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 364 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 365 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 366 | fun mk_exists v t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 367 | HOLogic.exists_const (fastype_of v) $ lambda v (incr_boundvars 1 t) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 368 | |
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 369 | fun pull_out_existential_constrs hol_ctxt t = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 370 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 371 | val k = maxidx_of_term t + 1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 372 | fun aux Ts num_exists t args seen = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 373 | case t of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 374 |         (t0 as Const (@{const_name Ex}, _)) $ Abs (s1, T1, t1) =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 375 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 376 | val (t1, seen') = aux (T1 :: Ts) (num_exists + 1) t1 [] [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 377 | val n = length seen' | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 378 | fun vars () = map2 (fresh_value_var Ts k n) (index_seq 0 n) seen' | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 379 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 380 | (equations_for_pulled_out_constrs HOLogic.mk_eq Ts k seen' | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 381 | |> List.foldl s_conj t1 |> fold mk_exists (vars ()) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 382 | |> curry3 Abs s1 T1 |> curry (op $) t0, seen) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 383 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 384 | | t1 $ t2 => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 385 | let val (t2, seen) = aux Ts num_exists t2 [] seen in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 386 | aux Ts num_exists t1 (t2 :: args) seen | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 387 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 388 | | Abs (s, T, t') => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 389 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 390 | val (t', seen) = aux (T :: Ts) 0 t' [] (map (incr_boundvars 1) seen) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 391 | in (list_comb (Abs (s, T, t'), args), map (incr_boundvars ~1) seen) end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 392 | | _ => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 393 | if num_exists > 0 then | 
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 394 | pull_out_constr_comb hol_ctxt Ts false k num_exists t args seen | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 395 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 396 | (list_comb (t, args), seen) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 397 | in aux [] 0 t [] [] |> fst end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 398 | |
| 55888 | 399 | fun destroy_pulled_out_constrs (hol_ctxt as {ctxt, ...}) axiom strong t =
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 400 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 401 | val num_occs_of_var = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 402 | fold_aterms (fn Var z => (fn f => fn z' => f z' |> z = z' ? Integer.add 1) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 403 | | _ => I) t (K 0) | 
| 56245 | 404 |     fun aux Ts careful ((t0 as Const (@{const_name Pure.eq}, _)) $ t1 $ t2) =
 | 
| 41793 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 405 | aux_eq Ts careful true t0 t1 t2 | 
| 56245 | 406 |       | aux Ts careful ((t0 as @{const Pure.imp}) $ t1 $ t2) =
 | 
| 41793 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 407 | t0 $ aux Ts false t1 $ aux Ts careful t2 | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 408 |       | aux Ts careful ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2) =
 | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 409 | aux_eq Ts careful true t0 t1 t2 | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 410 |       | aux Ts careful ((t0 as @{const HOL.implies}) $ t1 $ t2) =
 | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 411 | t0 $ aux Ts false t1 $ aux Ts careful t2 | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 412 | | aux Ts careful (Abs (s, T, t')) = Abs (s, T, aux (T :: Ts) careful t') | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 413 | | aux Ts careful (t1 $ t2) = aux Ts careful t1 $ aux Ts careful t2 | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 414 | | aux _ _ t = t | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 415 | and aux_eq Ts careful pass1 t0 t1 t2 = | 
| 41994 
c567c860caf6
always destroy constructor patterns, since this seems to be always useful
 blanchet parents: 
41876diff
changeset | 416 | ((if careful orelse | 
| 
c567c860caf6
always destroy constructor patterns, since this seems to be always useful
 blanchet parents: 
41876diff
changeset | 417 | not (strong orelse forall (is_constr_pattern ctxt) [t1, t2]) then | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 418 | raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 419 | else if axiom andalso is_Var t2 andalso | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 420 | num_occs_of_var (dest_Var t2) = 1 then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 421 |           @{const True}
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 422 | else case strip_comb t2 of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 423 | (* The first case is not as general as it could be. *) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 424 |           (Const (@{const_name PairBox}, _),
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 425 |                   [Const (@{const_name fst}, _) $ Var z1,
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 426 |                    Const (@{const_name snd}, _) $ Var z2]) =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 427 |           if z1 = z2 andalso num_occs_of_var z1 = 2 then @{const True}
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 428 | else raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 429 | | (Const (x as (s, T)), args) => | 
| 35386 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 430 | let | 
| 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 431 | val (arg_Ts, dataT) = strip_type T | 
| 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 432 | val n = length arg_Ts | 
| 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 433 | in | 
| 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 434 | if length args = n andalso | 
| 55888 | 435 |                (is_constr ctxt x orelse s = @{const_name Pair} orelse
 | 
| 35386 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 436 |                 x = (@{const_name Suc}, nat_T --> nat_T)) andalso
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 437 | (not careful orelse not (is_Var t1) orelse | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 438 | String.isPrefix val_var_prefix (fst (fst (dest_Var t1)))) then | 
| 41793 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 439 | s_let Ts "l" (n + 1) dataT bool_T | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 440 | (fn t1 => | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 441 | discriminate_value hol_ctxt x t1 :: | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 442 | map3 (sel_eq Ts x t1) (index_seq 0 n) arg_Ts args | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 443 | |> foldr1 s_conj) t1 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 444 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 445 | raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 446 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 447 | | _ => raise SAME ()) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 448 | |> body_type (type_of t0) = prop_T ? HOLogic.mk_Trueprop) | 
| 41793 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 449 | handle SAME () => if pass1 then aux_eq Ts careful false t0 t2 t1 | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 450 | else t0 $ aux Ts false t2 $ aux Ts false t1 | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 451 | and sel_eq Ts x t n nth_T nth_t = | 
| 55888 | 452 | HOLogic.eq_const nth_T $ nth_t $ select_nth_constr_arg ctxt x t n nth_T | 
| 41793 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 453 | |> aux Ts false | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 454 | in aux [] axiom t end | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 455 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 456 | (** Destruction of universal and existential equalities **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 457 | |
| 56245 | 458 | fun curry_assms (@{const Pure.imp} $ (@{const Trueprop}
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 459 |                                    $ (@{const HOL.conj} $ t1 $ t2)) $ t3) =
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 460 | curry_assms (Logic.list_implies ([t1, t2] |> map HOLogic.mk_Trueprop, t3)) | 
| 56245 | 461 |   | curry_assms (@{const Pure.imp} $ t1 $ t2) =
 | 
| 462 |     @{const Pure.imp} $ curry_assms t1 $ curry_assms t2
 | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 463 | | curry_assms t = t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 464 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 465 | val destroy_universal_equalities = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 466 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 467 | fun aux prems zs t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 468 | case t of | 
| 56245 | 469 |         @{const Pure.imp} $ t1 $ t2 => aux_implies prems zs t1 t2
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 470 | | _ => Logic.list_implies (rev prems, t) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 471 | and aux_implies prems zs t1 t2 = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 472 | case t1 of | 
| 56245 | 473 |         Const (@{const_name Pure.eq}, _) $ Var z $ t' => aux_eq prems zs z t' t1 t2
 | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 474 |       | @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ Var z $ t') =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 475 | aux_eq prems zs z t' t1 t2 | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 476 |       | @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ t' $ Var z) =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 477 | aux_eq prems zs z t' t1 t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 478 | | _ => aux (t1 :: prems) (Term.add_vars t1 zs) t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 479 | and aux_eq prems zs z t' t1 t2 = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 480 | if not (member (op =) zs z) andalso | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 481 | not (exists_subterm (curry (op =) (Var z)) t') then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 482 | aux prems zs (subst_free [(Var z, t')] t2) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 483 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 484 | aux (t1 :: prems) (Term.add_vars t1 zs) t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 485 | in aux [] [] end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 486 | |
| 55888 | 487 | fun find_bound_assign ctxt j = | 
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 488 | let | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 489 | fun do_term _ [] = NONE | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 490 | | do_term seen (t :: ts) = | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 491 | let | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 492 | fun do_eq pass1 t1 t2 = | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 493 | (if loose_bvar1 (t2, j) then | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 494 | if pass1 then do_eq false t2 t1 else raise SAME () | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 495 | else case t1 of | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 496 | Bound j' => if j' = j then SOME (t2, ts @ seen) else raise SAME () | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35408diff
changeset | 497 |              | Const (s, Type (@{type_name fun}, [T1, T2])) $ Bound j' =>
 | 
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 498 | if j' = j andalso | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 499 |                   s = nth_sel_name_for_constr_name @{const_name FunBox} 0 then
 | 
| 55888 | 500 | SOME (construct_value ctxt | 
| 37256 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 501 |                                        (@{const_name FunBox}, T2 --> T1) [t2],
 | 
| 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 502 | ts @ seen) | 
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 503 | else | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 504 | raise SAME () | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 505 | | _ => raise SAME ()) | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 506 | handle SAME () => do_term (t :: seen) ts | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 507 | in | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 508 | case t of | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 509 |             Const (@{const_name HOL.eq}, _) $ t1 $ t2 => do_eq true t1 t2
 | 
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 510 | | _ => do_term (t :: seen) ts | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 511 | end | 
| 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 512 | in do_term end | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 513 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 514 | fun subst_one_bound j arg t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 515 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 516 | fun aux (Bound i, lev) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 517 | if i < lev then raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 518 | else if i = lev then incr_boundvars (lev - j) arg | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 519 | else Bound (i - 1) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 520 | | aux (Abs (a, T, body), lev) = Abs (a, T, aux (body, lev + 1)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 521 | | aux (f $ t, lev) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 522 | (aux (f, lev) $ (aux (t, lev) handle SAME () => t) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 523 | handle SAME () => f $ aux (t, lev)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 524 | | aux _ = raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 525 | in aux (t, j) handle SAME () => t end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 526 | |
| 55888 | 527 | fun destroy_existential_equalities ({ctxt, ...} : hol_context) =
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 528 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 529 | fun kill [] [] ts = foldr1 s_conj ts | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 530 | | kill (s :: ss) (T :: Ts) ts = | 
| 55888 | 531 | (case find_bound_assign ctxt (length ss) [] ts of | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 532 |            SOME (_, []) => @{const True}
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 533 | | SOME (arg_t, ts) => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 534 | kill ss Ts (map (subst_one_bound (length ss) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 535 | (incr_bv (~1, length ss + 1, arg_t))) ts) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 536 | | NONE => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 537 |            Const (@{const_name Ex}, (T --> bool_T) --> bool_T)
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 538 | $ Abs (s, T, kill ss Ts ts)) | 
| 40722 
441260986b63
make two copies (!) of Library.UnequalLengths coincide with ListPair.UnequalLengths;
 wenzelm parents: 
39687diff
changeset | 539 | | kill _ _ _ = raise ListPair.UnequalLengths | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 540 |     fun gather ss Ts (Const (@{const_name Ex}, _) $ Abs (s1, T1, t1)) =
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 541 | gather (ss @ [s1]) (Ts @ [T1]) t1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 542 | | gather [] [] (Abs (s, T, t1)) = Abs (s, T, gather [] [] t1) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 543 | | gather [] [] (t1 $ t2) = gather [] [] t1 $ gather [] [] t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 544 | | gather [] [] t = t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 545 | | gather ss Ts t = kill ss Ts (conjuncts_of (gather [] [] t)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 546 | in gather [] [] end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 547 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 548 | (** Skolemization **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 549 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 550 | fun skolem_prefix_for k j = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 551 | skolem_prefix ^ string_of_int k ^ "@" ^ string_of_int j ^ name_sep | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 552 | |
| 41791 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 blanchet parents: 
41052diff
changeset | 553 | fun skolemize_term_and_more (hol_ctxt as {thy, def_tables, skolems, ...})
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 554 | skolem_depth = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 555 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 556 | val incrs = map (Integer.add 1) | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 557 | fun aux ss Ts js skolemizable polar t = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 558 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 559 | fun do_quantifier quant_s quant_T abs_s abs_T t = | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 560 | (if not (loose_bvar1 (t, 0)) then | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 561 | aux ss Ts js skolemizable polar (incr_boundvars ~1 t) | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 562 | else if is_positive_existential polar quant_s then | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 563 | let | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 564 | val j = length (!skolems) + 1 | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 565 | in | 
| 47754 
11b4395aaf0c
remove too aggressive skolemization optimization (prevented discovery of a model in SYN994^1)
 blanchet parents: 
47669diff
changeset | 566 | if skolemizable andalso length js <= skolem_depth then | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 567 | let | 
| 47754 
11b4395aaf0c
remove too aggressive skolemization optimization (prevented discovery of a model in SYN994^1)
 blanchet parents: 
47669diff
changeset | 568 | val sko_s = skolem_prefix_for (length js) j ^ abs_s | 
| 
11b4395aaf0c
remove too aggressive skolemization optimization (prevented discovery of a model in SYN994^1)
 blanchet parents: 
47669diff
changeset | 569 | val _ = Unsynchronized.change skolems (cons (sko_s, ss)) | 
| 
11b4395aaf0c
remove too aggressive skolemization optimization (prevented discovery of a model in SYN994^1)
 blanchet parents: 
47669diff
changeset | 570 | val sko_t = list_comb (Const (sko_s, rev Ts ---> abs_T), | 
| 
11b4395aaf0c
remove too aggressive skolemization optimization (prevented discovery of a model in SYN994^1)
 blanchet parents: 
47669diff
changeset | 571 | map Bound (rev js)) | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 572 | val abs_t = Abs (abs_s, abs_T, | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 573 | aux ss Ts (incrs js) skolemizable polar t) | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 574 | in | 
| 47754 
11b4395aaf0c
remove too aggressive skolemization optimization (prevented discovery of a model in SYN994^1)
 blanchet parents: 
47669diff
changeset | 575 | if null js then | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 576 | s_betapply Ts (abs_t, sko_t) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 577 | else | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 578 |                      Const (@{const_name Let}, abs_T --> quant_T) $ sko_t
 | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 579 | $ abs_t | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 580 | end | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 581 | else | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 582 | raise SAME () | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 583 | end | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 584 | else | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 585 | raise SAME ()) | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 586 | handle SAME () => | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 587 | Const (quant_s, quant_T) | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 588 | $ Abs (abs_s, abs_T, | 
| 38166 
28bb89672cc7
fix Skolemizer -- last week's optimizations resulted in UnequalLengths errors
 blanchet parents: 
38165diff
changeset | 589 | aux (abs_s :: ss) (abs_T :: Ts) (0 :: incrs js) | 
| 
28bb89672cc7
fix Skolemizer -- last week's optimizations resulted in UnequalLengths errors
 blanchet parents: 
38165diff
changeset | 590 | (skolemizable andalso | 
| 
28bb89672cc7
fix Skolemizer -- last week's optimizations resulted in UnequalLengths errors
 blanchet parents: 
38165diff
changeset | 591 | not (is_higher_order_type abs_T)) polar t) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 592 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 593 | case t of | 
| 56245 | 594 |           Const (s0 as @{const_name Pure.all}, T0) $ Abs (s1, T1, t1) =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 595 | do_quantifier s0 T0 s1 T1 t1 | 
| 56245 | 596 |         | @{const Pure.imp} $ t1 $ t2 =>
 | 
| 597 |           @{const Pure.imp} $ aux ss Ts js skolemizable (flip_polarity polar) t1
 | |
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 598 | $ aux ss Ts js skolemizable polar t2 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 599 |         | @{const Pure.conjunction} $ t1 $ t2 =>
 | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 600 |           @{const Pure.conjunction} $ aux ss Ts js skolemizable polar t1
 | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 601 | $ aux ss Ts js skolemizable polar t2 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 602 |         | @{const Trueprop} $ t1 =>
 | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 603 |           @{const Trueprop} $ aux ss Ts js skolemizable polar t1
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 604 |         | @{const Not} $ t1 =>
 | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 605 |           @{const Not} $ aux ss Ts js skolemizable (flip_polarity polar) t1
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 606 |         | Const (s0 as @{const_name All}, T0) $ Abs (s1, T1, t1) =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 607 | do_quantifier s0 T0 s1 T1 t1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 608 |         | Const (s0 as @{const_name Ex}, T0) $ Abs (s1, T1, t1) =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 609 | do_quantifier s0 T0 s1 T1 t1 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 610 |         | @{const HOL.conj} $ t1 $ t2 =>
 | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 611 | s_conj (pairself (aux ss Ts js skolemizable polar) (t1, t2)) | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 612 |         | @{const HOL.disj} $ t1 $ t2 =>
 | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 613 | s_disj (pairself (aux ss Ts js skolemizable polar) (t1, t2)) | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38240diff
changeset | 614 |         | @{const HOL.implies} $ t1 $ t2 =>
 | 
| 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38240diff
changeset | 615 |           @{const HOL.implies} $ aux ss Ts js skolemizable (flip_polarity polar) t1
 | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 616 | $ aux ss Ts js skolemizable polar t2 | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 617 |         | (t0 as Const (@{const_name Let}, _)) $ t1 $ t2 =>
 | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 618 | t0 $ t1 $ aux ss Ts js skolemizable polar t2 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 619 | | Const (x as (s, T)) => | 
| 55889 | 620 | if is_raw_inductive_pred hol_ctxt x andalso | 
| 621 | not (is_raw_equational_fun hol_ctxt x) andalso | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 622 | not (is_well_founded_inductive_pred hol_ctxt x) then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 623 | let | 
| 41791 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 blanchet parents: 
41052diff
changeset | 624 | val gfp = (fixpoint_kind_of_const thy def_tables x = Gfp) | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 625 | val (pref, connective) = | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 626 |                 if gfp then (lbfp_prefix, @{const HOL.disj})
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 627 |                 else (ubfp_prefix, @{const HOL.conj})
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 628 | fun pos () = unrolled_inductive_pred_const hol_ctxt gfp x | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 629 | |> aux ss Ts js skolemizable polar | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 630 | fun neg () = Const (pref ^ s, T) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 631 | in | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 632 | case polar |> gfp ? flip_polarity of | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 633 | Pos => pos () | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 634 | | Neg => neg () | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 635 | | Neut => | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 636 | let | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 637 | val arg_Ts = binder_types T | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 638 | fun app f = | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 639 | list_comb (f (), map Bound (length arg_Ts - 1 downto 0)) | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 640 | in | 
| 44241 | 641 | fold_rev absdummy arg_Ts (connective $ app pos $ app neg) | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 642 | end | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 643 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 644 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 645 | Const x | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 646 | | t1 $ t2 => | 
| 38166 
28bb89672cc7
fix Skolemizer -- last week's optimizations resulted in UnequalLengths errors
 blanchet parents: 
38165diff
changeset | 647 | s_betapply Ts (aux ss Ts js false polar t1, | 
| 
28bb89672cc7
fix Skolemizer -- last week's optimizations resulted in UnequalLengths errors
 blanchet parents: 
38165diff
changeset | 648 | aux ss Ts js false Neut t2) | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 649 | | Abs (s, T, t1) => | 
| 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 650 | Abs (s, T, aux ss Ts (incrs js) skolemizable polar t1) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 651 | | _ => t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 652 | end | 
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 653 | in aux [] [] [] true Pos end | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 654 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 655 | (** Function specialization **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 656 | |
| 56245 | 657 | fun params_in_equation (@{const Pure.imp} $ _ $ t2) = params_in_equation t2
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 658 |   | params_in_equation (@{const Trueprop} $ t1) = params_in_equation t1
 | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 659 |   | params_in_equation (Const (@{const_name HOL.eq}, _) $ t1 $ _) =
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 660 | snd (strip_comb t1) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 661 | | params_in_equation _ = [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 662 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 663 | fun specialize_fun_axiom x x' fixed_js fixed_args extra_args t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 664 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 665 | val k = fold Integer.max (map maxidx_of_term (fixed_args @ extra_args)) 0 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 666 | + 1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 667 | val t = map_aterms (fn Var ((s, i), T) => Var ((s, k + i), T) | t' => t') t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 668 | val fixed_params = filter_indices fixed_js (params_in_equation t) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 669 | fun aux args (Abs (s, T, t)) = list_comb (Abs (s, T, aux [] t), args) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 670 | | aux args (t1 $ t2) = aux (aux [] t2 :: args) t1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 671 | | aux args t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 672 | if t = Const x then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 673 | list_comb (Const x', extra_args @ filter_out_indices fixed_js args) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 674 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 675 | let val j = find_index (curry (op =) t) fixed_params in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 676 | list_comb (if j >= 0 then nth fixed_args j else t, args) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 677 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 678 | in aux [] t end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 679 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 680 | fun static_args_in_term ({ersatz_table, ...} : hol_context) x t =
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 681 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 682 | fun fun_calls (Abs (_, _, t)) _ = fun_calls t [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 683 | | fun_calls (t1 $ t2) args = fun_calls t2 [] #> fun_calls t1 (t2 :: args) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 684 | | fun_calls t args = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 685 | (case t of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 686 | Const (x' as (s', T')) => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 687 | x = x' orelse (case AList.lookup (op =) ersatz_table s' of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 688 | SOME s'' => x = (s'', T') | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 689 | | NONE => false) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 690 | | _ => false) ? cons args | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 691 | fun call_sets [] [] vs = [vs] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 692 | | call_sets [] uss vs = vs :: call_sets uss [] [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 693 | | call_sets ([] :: _) _ _ = [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 694 | | call_sets ((t :: ts) :: tss) uss vs = | 
| 39687 | 695 | Ord_List.insert Term_Ord.term_ord t vs |> call_sets tss (ts :: uss) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 696 | val sets = call_sets (fun_calls t [] []) [] [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 697 | val indexed_sets = sets ~~ (index_seq 0 (length sets)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 698 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 699 | fold_rev (fn (set, j) => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 700 | case set of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 701 | [Var _] => AList.lookup (op =) indexed_sets set = SOME j | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 702 | ? cons (j, NONE) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 703 | | [t as Const _] => cons (j, SOME t) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 704 | | [t as Free _] => cons (j, SOME t) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 705 | | _ => I) indexed_sets [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 706 | end | 
| 55889 | 707 | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 708 | fun static_args_in_terms hol_ctxt x = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 709 | map (static_args_in_term hol_ctxt x) | 
| 39687 | 710 | #> fold1 (Ord_List.inter (prod_ord int_ord (option_ord Term_Ord.term_ord))) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 711 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 712 | fun overlapping_indices [] _ = [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 713 | | overlapping_indices _ [] = [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 714 | | overlapping_indices (ps1 as (j1, t1) :: ps1') (ps2 as (j2, t2) :: ps2') = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 715 | if j1 < j2 then overlapping_indices ps1' ps2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 716 | else if j1 > j2 then overlapping_indices ps1 ps2' | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 717 | else overlapping_indices ps1' ps2' |> the_default t2 t1 = t2 ? cons j1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 718 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 719 | fun special_prefix_for j = special_prefix ^ string_of_int j ^ name_sep | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 720 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 721 | (* If a constant's definition is picked up deeper than this threshold, we | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 722 | prevent excessive specialization by not specializing it. *) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 723 | val special_max_depth = 20 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 724 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 725 | val bound_var_prefix = "b" | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 726 | |
| 38165 
9797be44df23
prevent generation of needless specialized constants etc.
 blanchet parents: 
37928diff
changeset | 727 | fun special_fun_aconv ((x1, js1, ts1), (x2, js2, ts2)) = | 
| 
9797be44df23
prevent generation of needless specialized constants etc.
 blanchet parents: 
37928diff
changeset | 728 | x1 = x2 andalso js1 = js2 andalso length ts1 = length ts2 andalso | 
| 
9797be44df23
prevent generation of needless specialized constants etc.
 blanchet parents: 
37928diff
changeset | 729 | forall (op aconv) (ts1 ~~ ts2) | 
| 
9797be44df23
prevent generation of needless specialized constants etc.
 blanchet parents: 
37928diff
changeset | 730 | |
| 38204 | 731 | fun specialize_consts_in_term | 
| 55888 | 732 |         (hol_ctxt as {ctxt, thy, specialize, def_tables, simp_table,
 | 
| 39359 | 733 | special_funs, ...}) def depth t = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 734 | if not specialize orelse depth > special_max_depth then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 735 | t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 736 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 737 | let | 
| 38206 | 738 | val blacklist = | 
| 739 | if def then case term_under_def t of Const x => [x] | _ => [] else [] | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 740 | fun aux args Ts (Const (x as (s, T))) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 741 | ((if not (member (op =) blacklist x) andalso not (null args) andalso | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 742 | not (String.isPrefix special_prefix s) andalso | 
| 55888 | 743 | not (is_built_in_const x) andalso | 
| 55889 | 744 | (is_equational_fun hol_ctxt x orelse | 
| 41791 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 blanchet parents: 
41052diff
changeset | 745 | (is_some (def_of_const thy def_tables x) andalso | 
| 38204 | 746 | not (is_of_class_const thy x) andalso | 
| 55888 | 747 | not (is_constr ctxt x) andalso | 
| 38204 | 748 | not (is_choice_spec_fun hol_ctxt x))) then | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 749 | let | 
| 41793 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 750 | val eligible_args = | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 751 | filter (is_special_eligible_arg true Ts o snd) | 
| 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 blanchet parents: 
41791diff
changeset | 752 | (index_seq 0 (length args) ~~ args) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 753 | val _ = not (null eligible_args) orelse raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 754 | val old_axs = equational_fun_axioms hol_ctxt x | 
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 755 | |> map (destroy_existential_equalities hol_ctxt) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 756 | val static_params = static_args_in_terms hol_ctxt x old_axs | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 757 | val fixed_js = overlapping_indices static_params eligible_args | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 758 | val _ = not (null fixed_js) orelse raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 759 | val fixed_args = filter_indices fixed_js args | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 760 | val vars = fold Term.add_vars fixed_args [] | 
| 35408 | 761 | |> sort (Term_Ord.fast_indexname_ord o pairself fst) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 762 | val bound_js = fold (fn t => fn js => add_loose_bnos (t, 0, js)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 763 | fixed_args [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 764 | |> sort int_ord | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 765 | val live_args = filter_out_indices fixed_js args | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 766 | val extra_args = map Var vars @ map Bound bound_js @ live_args | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 767 | val extra_Ts = map snd vars @ filter_indices bound_js Ts | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 768 | val k = maxidx_of_term t + 1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 769 | fun var_for_bound_no j = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 770 | Var ((bound_var_prefix ^ | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 771 | nat_subscript (find_index (curry (op =) j) bound_js | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 772 | + 1), k), | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 773 | nth Ts j) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 774 | val fixed_args_in_axiom = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 775 | map (curry subst_bounds | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 776 | (map var_for_bound_no (index_seq 0 (length Ts)))) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 777 | fixed_args | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 778 | in | 
| 38165 
9797be44df23
prevent generation of needless specialized constants etc.
 blanchet parents: 
37928diff
changeset | 779 | case AList.lookup special_fun_aconv (!special_funs) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 780 | (x, fixed_js, fixed_args_in_axiom) of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 781 | SOME x' => list_comb (Const x', extra_args) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 782 | | NONE => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 783 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 784 | val extra_args_in_axiom = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 785 | map Var vars @ map var_for_bound_no bound_js | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 786 | val x' as (s', _) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 787 | (special_prefix_for (length (!special_funs) + 1) ^ s, | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 788 | extra_Ts @ filter_out_indices fixed_js (binder_types T) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 789 | ---> body_type T) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 790 | val new_axs = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 791 | map (specialize_fun_axiom x x' fixed_js | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 792 | fixed_args_in_axiom extra_args_in_axiom) old_axs | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 793 | val _ = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 794 | Unsynchronized.change special_funs | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 795 | (cons ((x, fixed_js, fixed_args_in_axiom), x')) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 796 | val _ = add_simps simp_table s' new_axs | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 797 | in list_comb (Const x', extra_args) end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 798 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 799 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 800 | raise SAME ()) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 801 | handle SAME () => list_comb (Const x, args)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 802 | | aux args Ts (Abs (s, T, t)) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 803 | list_comb (Abs (s, T, aux [] (T :: Ts) t), args) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 804 | | aux args Ts (t1 $ t2) = aux (aux [] Ts t2 :: args) Ts t1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 805 | | aux args _ t = list_comb (t, args) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 806 | in aux [] [] t end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 807 | |
| 55889 | 808 | type special_triple = int list * term list * (string * typ) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 809 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 810 | val cong_var_prefix = "c" | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 811 | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 812 | fun special_congruence_axiom T (js1, ts1, x1) (js2, ts2, x2) = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 813 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 814 | val (bounds1, bounds2) = pairself (map Var o special_bounds) (ts1, ts2) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 815 | val Ts = binder_types T | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 816 | val max_j = fold (fold Integer.max) [js1, js2] ~1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 817 | val (eqs, (args1, args2)) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 818 | fold (fn j => case pairself (fn ps => AList.lookup (op =) ps j) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 819 | (js1 ~~ ts1, js2 ~~ ts2) of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 820 | (SOME t1, SOME t2) => apfst (cons (t1, t2)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 821 | | (SOME t1, NONE) => apsnd (apsnd (cons t1)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 822 | | (NONE, SOME t2) => apsnd (apfst (cons t2)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 823 | | (NONE, NONE) => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 824 | let val v = Var ((cong_var_prefix ^ nat_subscript j, 0), | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 825 | nth Ts j) in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 826 | apsnd (pairself (cons v)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 827 | end) (max_j downto 0) ([], ([], [])) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 828 | in | 
| 38165 
9797be44df23
prevent generation of needless specialized constants etc.
 blanchet parents: 
37928diff
changeset | 829 | Logic.list_implies (eqs |> filter_out (op aconv) |> distinct (op =) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 830 | |> map Logic.mk_equals, | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 831 | Logic.mk_equals (list_comb (Const x1, bounds1 @ args1), | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 832 | list_comb (Const x2, bounds2 @ args2))) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 833 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 834 | |
| 38170 | 835 | fun special_congruence_axioms (hol_ctxt as {special_funs, ...}) ts =
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 836 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 837 | val groups = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 838 | !special_funs | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 839 | |> map (fn ((x, js, ts), x') => (x, (js, ts, x'))) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 840 | |> AList.group (op =) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 841 | |> filter_out (is_equational_fun_surely_complete hol_ctxt o fst) | 
| 38170 | 842 | |> map (fn (x, zs) => | 
| 843 | (x, zs |> member (op =) ts (Const x) ? cons ([], [], x))) | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 844 | fun generality (js, _, _) = ~(length js) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 845 | fun is_more_specific (j1, t1, x1) (j2, t2, x2) = | 
| 38165 
9797be44df23
prevent generation of needless specialized constants etc.
 blanchet parents: 
37928diff
changeset | 846 | x1 <> x2 andalso length j2 < length j1 andalso | 
| 39687 | 847 | Ord_List.subset (prod_ord int_ord Term_Ord.term_ord) (j2 ~~ t2, j1 ~~ t1) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 848 | fun do_pass_1 _ [] [_] [_] = I | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 849 | | do_pass_1 T skipped _ [] = do_pass_2 T skipped | 
| 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 850 | | do_pass_1 T skipped all (z :: zs) = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 851 | case filter (is_more_specific z) all | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 852 | |> sort (int_ord o pairself generality) of | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 853 | [] => do_pass_1 T (z :: skipped) all zs | 
| 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 854 | | (z' :: _) => cons (special_congruence_axiom T z z') | 
| 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 855 | #> do_pass_1 T skipped all zs | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 856 | and do_pass_2 _ [] = I | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 857 | | do_pass_2 T (z :: zs) = | 
| 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 858 | fold (cons o special_congruence_axiom T z) zs #> do_pass_2 T zs | 
| 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 859 | in fold (fn ((_, T), zs) => do_pass_1 T [] zs zs) groups [] end | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 860 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 861 | (** Axiom selection **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 862 | |
| 38169 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 863 | fun defined_free_by_assumption t = | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 864 | let | 
| 47669 | 865 | fun do_equals u def = | 
| 866 | if exists_subterm (curry (op aconv) u) def then NONE else SOME u | |
| 38169 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 867 | in | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 868 | case t of | 
| 56245 | 869 |       Const (@{const_name Pure.eq}, _) $ (u as Free _) $ def => do_equals u def
 | 
| 47669 | 870 |     | @{const Trueprop}
 | 
| 871 |       $ (Const (@{const_name HOL.eq}, _) $ (u as Free _) $ def) =>
 | |
| 872 | do_equals u def | |
| 38169 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 873 | | _ => NONE | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 874 | end | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 875 | |
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 876 | fun assumption_exclusively_defines_free assm_ts t = | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 877 | case defined_free_by_assumption t of | 
| 47669 | 878 | SOME u => | 
| 879 | length (filter ((fn SOME u' => u aconv u' | NONE => false) | |
| 38169 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 880 | o defined_free_by_assumption) assm_ts) = 1 | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 881 | | NONE => false | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 882 | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 883 | fun all_table_entries table = Symtab.fold (append o snd) table [] | 
| 55889 | 884 | |
| 41791 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 blanchet parents: 
41052diff
changeset | 885 | fun extra_table tables s = | 
| 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 blanchet parents: 
41052diff
changeset | 886 | Symtab.make [(s, pairself all_table_entries tables |> op @)] | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 887 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 888 | fun eval_axiom_for_term j t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 889 | Logic.mk_equals (Const (eval_prefix ^ string_of_int j, fastype_of t), t) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 890 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 891 | val is_trivial_equation = the_default false o try (op aconv o Logic.dest_equals) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 892 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 893 | (* Prevents divergence in case of cyclic or infinite axiom dependencies. *) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 894 | val axioms_max_depth = 255 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 895 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 896 | fun axioms_for_term | 
| 55888 | 897 |         (hol_ctxt as {thy, ctxt, max_bisim_depth, user_axioms, evals,
 | 
| 898 | def_tables, nondef_table, choice_spec_table, nondefs, | |
| 899 | ...}) assm_ts neg_t = | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 900 | let | 
| 38170 | 901 | val (def_assm_ts, nondef_assm_ts) = | 
| 902 | List.partition (assumption_exclusively_defines_free assm_ts) assm_ts | |
| 903 | val def_assm_table = map (`(the o defined_free_by_assumption)) def_assm_ts | |
| 55889 | 904 | type accumulator = (string * typ) list * (term list * term list) | 
| 38206 | 905 | fun add_axiom get app def depth t (accum as (seen, axs)) = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 906 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 907 | val t = t |> unfold_defs_in_term hol_ctxt | 
| 38206 | 908 | |> skolemize_term_and_more hol_ctxt ~1 (* FIXME: why ~1? *) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 909 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 910 | if is_trivial_equation t then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 911 | accum | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 912 | else | 
| 38206 | 913 | let val t' = t |> specialize_consts_in_term hol_ctxt def depth in | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 914 | if exists (member (op aconv) (get axs)) [t, t'] then accum | 
| 38170 | 915 | else add_axioms_for_term (depth + 1) t' (seen, app (cons t') axs) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 916 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 917 | end | 
| 38206 | 918 | and add_def_axiom depth = add_axiom fst apfst true depth | 
| 919 | and add_nondef_axiom depth = add_axiom snd apsnd false depth | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 920 | and add_maybe_def_axiom depth t = | 
| 56245 | 921 |       (if head_of t <> @{const Pure.imp} then add_def_axiom
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 922 | else add_nondef_axiom) depth t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 923 | and add_eq_axiom depth t = | 
| 37256 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 924 | (if is_constr_pattern_formula ctxt t then add_def_axiom | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 925 | else add_nondef_axiom) depth t | 
| 38170 | 926 | and add_axioms_for_term depth t (accum as (seen, axs)) = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 927 | case t of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 928 | t1 $ t2 => accum |> fold (add_axioms_for_term depth) [t1, t2] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 929 | | Const (x as (s, T)) => | 
| 55888 | 930 | (if member (op aconv) seen t orelse is_built_in_const x then | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 931 | accum | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 932 | else | 
| 38170 | 933 | let val accum = (t :: seen, axs) in | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 934 | if depth > axioms_max_depth then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 935 |                raise TOO_LARGE ("Nitpick_Preproc.axioms_for_term.\
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 936 | \add_axioms_for_term", | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 937 |                                 "too many nested axioms (" ^
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 938 | string_of_int depth ^ ")") | 
| 37260 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37256diff
changeset | 939 | else if is_of_class_const thy x then | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 940 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 941 | val class = Logic.class_of_const s | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 942 |                  val of_class = Logic.mk_of_class (TVar (("'a", 0), [class]),
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 943 | class) | 
| 36555 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36389diff
changeset | 944 | val ax1 = try (specialize_type thy x) of_class | 
| 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36389diff
changeset | 945 | val ax2 = Option.map (specialize_type thy x o snd) | 
| 37260 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 blanchet parents: 
37256diff
changeset | 946 | (get_class_def thy class) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 947 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 948 | fold (add_maybe_def_axiom depth) (map_filter I [ax1, ax2]) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 949 | accum | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 950 | end | 
| 55888 | 951 | else if is_constr ctxt x then | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 952 | accum | 
| 37271 | 953 | else if is_descr (original_name s) then | 
| 954 | fold (add_nondef_axiom depth) (equational_fun_axioms hol_ctxt x) | |
| 955 | accum | |
| 55889 | 956 | else if is_equational_fun hol_ctxt x then | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 957 | fold (add_eq_axiom depth) (equational_fun_axioms hol_ctxt x) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 958 | accum | 
| 35807 
e4d1b5cbd429
added support for "specification" and "ax_specification" constructs to Nitpick
 blanchet parents: 
35718diff
changeset | 959 | else if is_choice_spec_fun hol_ctxt x then | 
| 
e4d1b5cbd429
added support for "specification" and "ax_specification" constructs to Nitpick
 blanchet parents: 
35718diff
changeset | 960 | fold (add_nondef_axiom depth) | 
| 
e4d1b5cbd429
added support for "specification" and "ax_specification" constructs to Nitpick
 blanchet parents: 
35718diff
changeset | 961 | (nondef_props_for_const thy true choice_spec_table x) accum | 
| 37256 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 962 | else if is_abs_fun ctxt x then | 
| 38215 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 963 | accum |> fold (add_nondef_axiom depth) | 
| 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 964 | (nondef_props_for_const thy false nondef_table x) | 
| 38240 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 blanchet parents: 
38215diff
changeset | 965 | |> (is_funky_typedef ctxt (range_type T) orelse | 
| 38215 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 966 | range_type T = nat_T) | 
| 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 967 | ? fold (add_maybe_def_axiom depth) | 
| 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 968 | (nondef_props_for_const thy true | 
| 41791 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 blanchet parents: 
41052diff
changeset | 969 | (extra_table def_tables s) x) | 
| 37256 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 970 | else if is_rep_fun ctxt x then | 
| 38215 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 971 | accum |> fold (add_nondef_axiom depth) | 
| 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 972 | (nondef_props_for_const thy false nondef_table x) | 
| 38240 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 blanchet parents: 
38215diff
changeset | 973 | |> (is_funky_typedef ctxt (range_type T) orelse | 
| 38215 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 974 | range_type T = nat_T) | 
| 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 975 | ? fold (add_maybe_def_axiom depth) | 
| 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 976 | (nondef_props_for_const thy true | 
| 41791 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 blanchet parents: 
41052diff
changeset | 977 | (extra_table def_tables s) x) | 
| 38215 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 978 | |> add_axioms_for_term depth | 
| 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 979 | (Const (mate_of_rep_fun ctxt x)) | 
| 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 980 | |> fold (add_def_axiom depth) | 
| 
1c7d7eaebdf2
quotient types registered as codatatypes are no longer quotient types
 blanchet parents: 
38207diff
changeset | 981 | (inverse_axioms_for_rep_fun ctxt x) | 
| 56243 | 982 |              else if s = @{const_name Pure.type} then
 | 
| 37253 | 983 | accum | 
| 41791 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 blanchet parents: 
41052diff
changeset | 984 | else case def_of_const thy def_tables x of | 
| 39345 | 985 | SOME _ => | 
| 38202 
379fb08da97b
prevent the expansion of too large definitions -- use equations for these instead
 blanchet parents: 
38190diff
changeset | 986 | fold (add_eq_axiom depth) (equational_fun_axioms hol_ctxt x) | 
| 
379fb08da97b
prevent the expansion of too large definitions -- use equations for these instead
 blanchet parents: 
38190diff
changeset | 987 | accum | 
| 
379fb08da97b
prevent the expansion of too large definitions -- use equations for these instead
 blanchet parents: 
38190diff
changeset | 988 | | NONE => | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 989 | accum |> user_axioms <> SOME false | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 990 | ? fold (add_nondef_axiom depth) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 991 | (nondef_props_for_const thy false nondef_table x) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 992 | end) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 993 | |> add_axioms_for_type depth T | 
| 47669 | 994 | | Free (_, T) => | 
| 38170 | 995 | (if member (op aconv) seen t then | 
| 996 | accum | |
| 47669 | 997 | else case AList.lookup (op =) def_assm_table t of | 
| 38170 | 998 | SOME t => add_def_axiom depth t accum | 
| 999 | | NONE => accum) | |
| 1000 | |> add_axioms_for_type depth T | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1001 | | Var (_, T) => add_axioms_for_type depth T accum | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1002 | | Bound _ => accum | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1003 | | Abs (_, T, t) => accum |> add_axioms_for_term depth t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1004 | |> add_axioms_for_type depth T | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1005 | and add_axioms_for_type depth T = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1006 | case T of | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35408diff
changeset | 1007 |         Type (@{type_name fun}, Ts) => fold (add_axioms_for_type depth) Ts
 | 
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38170diff
changeset | 1008 |       | Type (@{type_name prod}, Ts) => fold (add_axioms_for_type depth) Ts
 | 
| 46115 | 1009 |       | Type (@{type_name set}, Ts) => fold (add_axioms_for_type depth) Ts
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1010 |       | @{typ prop} => I
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1011 |       | @{typ bool} => I
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1012 | | TFree (_, S) => add_axioms_for_sort depth T S | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1013 | | TVar (_, S) => add_axioms_for_sort depth T S | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1014 | | Type (z as (_, Ts)) => | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1015 | fold (add_axioms_for_type depth) Ts | 
| 37256 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 1016 | #> (if is_pure_typedef ctxt T then | 
| 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 1017 | fold (add_maybe_def_axiom depth) (optimized_typedef_axioms ctxt z) | 
| 38240 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 blanchet parents: 
38215diff
changeset | 1018 | else if is_quot_type ctxt T then | 
| 55888 | 1019 | fold (add_def_axiom depth) (optimized_quot_type_axioms ctxt z) | 
| 38240 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 blanchet parents: 
38215diff
changeset | 1020 | else if max_bisim_depth >= 0 andalso is_codatatype ctxt T then | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1021 | fold (add_maybe_def_axiom depth) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1022 | (codatatype_bisim_axioms hol_ctxt T) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1023 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1024 | I) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1025 | and add_axioms_for_sort depth T S = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1026 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1027 | val supers = Sign.complete_sort thy S | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1028 | val class_axioms = | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
47909diff
changeset | 1029 | maps (fn class => map prop_of (Axclass.get_info thy class |> #axioms | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1030 | handle ERROR _ => [])) supers | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1031 | val monomorphic_class_axioms = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1032 | map (fn t => case Term.add_tvars t [] of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1033 | [] => t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1034 | | [(x, S)] => | 
| 36555 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36389diff
changeset | 1035 | monomorphic_term (Vartab.make [(x, (S, T))]) t | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1036 |                        | _ => raise TERM ("Nitpick_Preproc.axioms_for_term.\
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1037 | \add_axioms_for_sort", [t])) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1038 | class_axioms | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1039 | in fold (add_nondef_axiom depth) monomorphic_class_axioms end | 
| 42415 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 blanchet parents: 
42361diff
changeset | 1040 | val (mono_nondefs, poly_nondefs) = | 
| 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 blanchet parents: 
42361diff
changeset | 1041 | List.partition (null o Term.hidden_polymorphism) nondefs | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1042 | val eval_axioms = map2 eval_axiom_for_term (index_seq 0 (length evals)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1043 | evals | 
| 38170 | 1044 | val (seen, (defs, nondefs)) = | 
| 38169 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 1045 | ([], ([], [])) | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 1046 | |> add_axioms_for_term 1 neg_t | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 1047 | |> fold_rev (add_nondef_axiom 1) nondef_assm_ts | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 1048 | |> fold_rev (add_def_axiom 1) eval_axioms | 
| 42415 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 blanchet parents: 
42361diff
changeset | 1049 | |> user_axioms = SOME true ? fold (add_nondef_axiom 1) mono_nondefs | 
| 38170 | 1050 | val defs = defs @ special_congruence_axioms hol_ctxt seen | 
| 35386 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 1051 | val got_all_mono_user_axioms = | 
| 42415 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 blanchet parents: 
42361diff
changeset | 1052 | (user_axioms = SOME true orelse null mono_nondefs) | 
| 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 blanchet parents: 
42361diff
changeset | 1053 | in (neg_t :: nondefs, defs, got_all_mono_user_axioms, null poly_nondefs) end | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1054 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1055 | (** Simplification of constructor/selector terms **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1056 | |
| 37256 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 1057 | fun simplify_constrs_and_sels ctxt t = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1058 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1059 | fun is_nth_sel_on t' n (Const (s, _) $ t) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1060 | (t = t' andalso is_sel_like_and_no_discr s andalso | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1061 | sel_no_from_name s = n) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1062 | | is_nth_sel_on _ _ _ = false | 
| 47909 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 blanchet parents: 
47754diff
changeset | 1063 |     fun do_term (Const (@{const_name Nitpick.Rep_Frac}, _)
 | 
| 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 blanchet parents: 
47754diff
changeset | 1064 |                  $ (Const (@{const_name Nitpick.Abs_Frac}, _) $ t1)) [] =
 | 
| 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 blanchet parents: 
47754diff
changeset | 1065 | do_term t1 [] | 
| 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 blanchet parents: 
47754diff
changeset | 1066 |       | do_term (Const (@{const_name Nitpick.Abs_Frac}, _)
 | 
| 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 blanchet parents: 
47754diff
changeset | 1067 |                  $ (Const (@{const_name Nitpick.Rep_Frac}, _) $ t1)) [] =
 | 
| 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 blanchet parents: 
47754diff
changeset | 1068 | do_term t1 [] | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1069 | | do_term (t1 $ t2) args = do_term t1 (do_term t2 [] :: args) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1070 | | do_term (t as Const (x as (s, T))) (args as _ :: _) = | 
| 55889 | 1071 | ((if is_nonfree_constr ctxt x then | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1072 | if length args = num_binder_types T then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1073 | case hd args of | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1074 | Const (_, T') $ t' => | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1075 | if domain_type T' = body_type T andalso | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1076 | forall (uncurry (is_nth_sel_on t')) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1077 | (index_seq 0 (length args) ~~ args) then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1078 | t' | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1079 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1080 | raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1081 | | _ => raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1082 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1083 | raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1084 | else if is_sel_like_and_no_discr s then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1085 | case strip_comb (hd args) of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1086 | (Const (x' as (s', T')), ts') => | 
| 55889 | 1087 | if is_free_constr ctxt x' andalso | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1088 | constr_name_for_sel_like s = s' andalso | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1089 | not (exists is_pair_type (binder_types T')) then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1090 | list_comb (nth ts' (sel_no_from_name s), tl args) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1091 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1092 | raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1093 | | _ => raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1094 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1095 | raise SAME ()) | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 1096 | handle SAME () => s_betapplys [] (t, args)) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1097 | | do_term (Abs (s, T, t')) args = | 
| 37476 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 1098 | s_betapplys [] (Abs (s, T, do_term t' []), args) | 
| 
0681e46b4022
optimized code generated for datatype cases + more;
 blanchet parents: 
37271diff
changeset | 1099 | | do_term t args = s_betapplys [] (t, args) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1100 | in do_term t [] end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1101 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1102 | (** Quantifier massaging: Distributing quantifiers **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1103 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1104 | fun distribute_quantifiers t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1105 | case t of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1106 |     (t0 as Const (@{const_name All}, T0)) $ Abs (s, T1, t1) =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1107 | (case t1 of | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 1108 |        (t10 as @{const HOL.conj}) $ t11 $ t12 =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1109 | t10 $ distribute_quantifiers (t0 $ Abs (s, T1, t11)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1110 | $ distribute_quantifiers (t0 $ Abs (s, T1, t12)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1111 |      | (t10 as @{const Not}) $ t11 =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1112 |        t10 $ distribute_quantifiers (Const (@{const_name Ex}, T0)
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1113 | $ Abs (s, T1, t11)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1114 | | t1 => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1115 | if not (loose_bvar1 (t1, 0)) then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1116 | distribute_quantifiers (incr_boundvars ~1 t1) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1117 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1118 | t0 $ Abs (s, T1, distribute_quantifiers t1)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1119 |   | (t0 as Const (@{const_name Ex}, T0)) $ Abs (s, T1, t1) =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1120 | (case distribute_quantifiers t1 of | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 1121 |        (t10 as @{const HOL.disj}) $ t11 $ t12 =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1122 | t10 $ distribute_quantifiers (t0 $ Abs (s, T1, t11)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1123 | $ distribute_quantifiers (t0 $ Abs (s, T1, t12)) | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38240diff
changeset | 1124 |      | (t10 as @{const HOL.implies}) $ t11 $ t12 =>
 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1125 |        t10 $ distribute_quantifiers (Const (@{const_name All}, T0)
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1126 | $ Abs (s, T1, t11)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1127 | $ distribute_quantifiers (t0 $ Abs (s, T1, t12)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1128 |      | (t10 as @{const Not}) $ t11 =>
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1129 |        t10 $ distribute_quantifiers (Const (@{const_name All}, T0)
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1130 | $ Abs (s, T1, t11)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1131 | | t1 => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1132 | if not (loose_bvar1 (t1, 0)) then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1133 | distribute_quantifiers (incr_boundvars ~1 t1) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1134 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1135 | t0 $ Abs (s, T1, distribute_quantifiers t1)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1136 | | t1 $ t2 => distribute_quantifiers t1 $ distribute_quantifiers t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1137 | | Abs (s, T, t') => Abs (s, T, distribute_quantifiers t') | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1138 | | _ => t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1139 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1140 | (** Quantifier massaging: Pushing quantifiers inward **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1141 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1142 | fun renumber_bounds j n f t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1143 | case t of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1144 | t1 $ t2 => renumber_bounds j n f t1 $ renumber_bounds j n f t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1145 | | Abs (s, T, t') => Abs (s, T, renumber_bounds (j + 1) n f t') | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1146 | | Bound j' => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1147 | Bound (if j' >= j andalso j' < j + n then f (j' - j) + j else j') | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1148 | | _ => t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1149 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1150 | (* Maximum number of quantifiers in a cluster for which the exponential | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1151 | algorithm is used. Larger clusters use a heuristic inspired by Claessen & | 
| 35386 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 1152 | Soerensson's polynomial binary splitting procedure (p. 5 of their MODEL 2003 | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1153 | paper). *) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1154 | val quantifier_cluster_threshold = 7 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1155 | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1156 | val push_quantifiers_inward = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1157 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1158 | fun aux quant_s ss Ts t = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1159 | (case t of | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1160 | Const (s0, _) $ Abs (s1, T1, t1 as _ $ _) => | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1161 | if s0 = quant_s then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1162 | aux s0 (s1 :: ss) (T1 :: Ts) t1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1163 | else if quant_s = "" andalso | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1164 |                  (s0 = @{const_name All} orelse s0 = @{const_name Ex}) then
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1165 | aux s0 [s1] [T1] t1 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1166 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1167 | raise SAME () | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1168 | | _ => raise SAME ()) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1169 | handle SAME () => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1170 | case t of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1171 | t1 $ t2 => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1172 | if quant_s = "" then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1173 | aux "" [] [] t1 $ aux "" [] [] t2 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1174 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1175 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1176 | fun big_union proj ps = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1177 | fold (fold (insert (op =)) o proj) ps [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1178 | val (ts, connective) = strip_any_connective t | 
| 41860 
49d0fc8c418c
improved "is_special_eligible_arg" further, by approximating the cardinality of the argument (as is done quite successfully elsewhere)
 blanchet parents: 
41803diff
changeset | 1179 | val T_costs = map typical_card_of_type Ts | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1180 | val t_costs = map size_of_term ts | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1181 | val num_Ts = length Ts | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1182 | val flip = curry (op -) (num_Ts - 1) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1183 | val t_boundss = map (map flip o loose_bnos) ts | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1184 | fun merge costly_boundss [] = costly_boundss | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1185 | | merge costly_boundss (j :: js) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1186 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1187 | val (yeas, nays) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1188 | List.partition (fn (bounds, _) => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1189 | member (op =) bounds j) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1190 | costly_boundss | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1191 | val yeas_bounds = big_union fst yeas | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1192 | val yeas_cost = Integer.sum (map snd yeas) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1193 | * nth T_costs j | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1194 | in merge ((yeas_bounds, yeas_cost) :: nays) js end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1195 | val cost = Integer.sum o map snd oo merge | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1196 | fun heuristically_best_permutation _ [] = [] | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1197 | | heuristically_best_permutation costly_boundss js = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1198 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1199 | val (costly_boundss, (j, js)) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1200 | js |> map (`(merge costly_boundss o single)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1201 | |> sort (int_ord | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1202 | o pairself (Integer.sum o map snd o fst)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1203 | |> split_list |>> hd ||> pairf hd tl | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1204 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1205 | j :: heuristically_best_permutation costly_boundss js | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1206 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1207 | val js = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1208 | if length Ts <= quantifier_cluster_threshold then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1209 | all_permutations (index_seq 0 num_Ts) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1210 | |> map (`(cost (t_boundss ~~ t_costs))) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1211 | |> sort (int_ord o pairself fst) |> hd |> snd | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1212 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1213 | heuristically_best_permutation (t_boundss ~~ t_costs) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1214 | (index_seq 0 num_Ts) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1215 | val back_js = map (fn j => find_index (curry (op =) j) js) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1216 | (index_seq 0 num_Ts) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1217 | val ts = map (renumber_bounds 0 num_Ts (nth back_js o flip)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1218 | ts | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1219 | fun mk_connection [] = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1220 |                        raise ARG ("Nitpick_Preproc.push_quantifiers_inward.aux.\
 | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1221 | \mk_connection", "") | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1222 | | mk_connection ts_cum_bounds = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1223 | ts_cum_bounds |> map fst | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1224 | |> foldr1 (fn (t1, t2) => connective $ t1 $ t2) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1225 | fun build ts_cum_bounds [] = ts_cum_bounds |> mk_connection | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1226 | | build ts_cum_bounds (j :: js) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1227 | let | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1228 | val (yeas, nays) = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1229 | List.partition (fn (_, bounds) => | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1230 | member (op =) bounds j) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1231 | ts_cum_bounds | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1232 | ||> map (apfst (incr_boundvars ~1)) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1233 | in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1234 | if null yeas then | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1235 | build nays js | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1236 | else | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1237 | let val T = nth Ts (flip j) in | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1238 | build ((Const (quant_s, (T --> bool_T) --> bool_T) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1239 | $ Abs (nth ss (flip j), T, | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1240 | mk_connection yeas), | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1241 | big_union snd yeas) :: nays) js | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1242 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1243 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1244 | in build (ts ~~ t_boundss) js end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1245 | | Abs (s, T, t') => Abs (s, T, aux "" [] [] t') | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1246 | | _ => t | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1247 | in aux "" [] [] end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1248 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1249 | (** Preprocessor entry point **) | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1250 | |
| 37928 
24785fa2416c
do a better job at Skolemizing in Nitpick, for TPTP FOF
 blanchet parents: 
37704diff
changeset | 1251 | val max_skolem_depth = 3 | 
| 36389 
8228b3a4a2ba
remove "skolemize" option from Nitpick, since Skolemization is always useful
 blanchet parents: 
36388diff
changeset | 1252 | |
| 38169 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 1253 | fun preprocess_formulas | 
| 55888 | 1254 |         (hol_ctxt as {ctxt, binary_ints, destroy_constrs, boxes, needs, ...})
 | 
| 1255 | assm_ts neg_t = | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1256 | let | 
| 35386 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 1257 | val (nondef_ts, def_ts, got_all_mono_user_axioms, no_poly_user_axioms) = | 
| 38169 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 1258 | neg_t |> unfold_defs_in_term hol_ctxt | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 1259 | |> close_form | 
| 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 1260 | |> skolemize_term_and_more hol_ctxt max_skolem_depth | 
| 38206 | 1261 | |> specialize_consts_in_term hol_ctxt false 0 | 
| 38169 
b51784515471
optimize local "def"s by treating them as definitions
 blanchet parents: 
38166diff
changeset | 1262 | |> axioms_for_term hol_ctxt assm_ts | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1263 | val binarize = | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1264 | case binary_ints of | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1265 | SOME false => false | 
| 35718 | 1266 | | _ => forall (may_use_binary_ints false) nondef_ts andalso | 
| 1267 | forall (may_use_binary_ints true) def_ts andalso | |
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 1268 | (binary_ints = SOME true orelse | 
| 35386 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 1269 | exists should_use_binary_ints (nondef_ts @ def_ts)) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1270 | val box = exists (not_equal (SOME false) o snd) boxes | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1271 | val table = | 
| 35386 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 1272 | Termtab.empty | 
| 37256 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 1273 | |> box ? fold (add_to_uncurry_table ctxt) (nondef_ts @ def_ts) | 
| 41803 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41793diff
changeset | 1274 | fun do_middle def = | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1275 | binarize ? binarize_nat_and_int_in_term | 
| 36388 | 1276 | #> box ? uncurry_term table | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1277 | #> box ? box_fun_and_pair_in_term hol_ctxt def | 
| 41803 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41793diff
changeset | 1278 | fun do_tail def = | 
| 46102 | 1279 | destroy_set_Collect | 
| 1280 | #> destroy_constrs ? (pull_out_universal_constrs hol_ctxt def | |
| 1281 | #> pull_out_existential_constrs hol_ctxt) | |
| 41994 
c567c860caf6
always destroy constructor patterns, since this seems to be always useful
 blanchet parents: 
41876diff
changeset | 1282 | #> destroy_pulled_out_constrs hol_ctxt def destroy_constrs | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1283 | #> curry_assms | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1284 | #> destroy_universal_equalities | 
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 1285 | #> destroy_existential_equalities hol_ctxt | 
| 37256 
0dca1ec52999
thread along context instead of theory for typedef lookup
 blanchet parents: 
37253diff
changeset | 1286 | #> simplify_constrs_and_sels ctxt | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1287 | #> distribute_quantifiers | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1288 | #> push_quantifiers_inward | 
| 35078 
6fd1052fe463
optimization to quantifiers in Nitpick's handling of simp rules + renamed some SAT solvers
 blanchet parents: 
35070diff
changeset | 1289 | #> close_form | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1290 | #> Term.map_abs_vars shortest_name | 
| 41803 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41793diff
changeset | 1291 | val nondef_ts = nondef_ts |> map (do_middle false) | 
| 41875 | 1292 | val need_ts = | 
| 41876 | 1293 | case needs of | 
| 1294 | SOME needs => | |
| 1295 | needs |> map (unfold_defs_in_term hol_ctxt #> do_middle false) | |
| 1296 | | NONE => [] (* FIXME: Implement inference. *) | |
| 41803 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41793diff
changeset | 1297 | val nondef_ts = nondef_ts |> map (do_tail false) | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41793diff
changeset | 1298 | val def_ts = def_ts |> map (do_middle true #> do_tail true) | 
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1299 | in | 
| 41875 | 1300 | (nondef_ts, def_ts, need_ts, got_all_mono_user_axioms, no_poly_user_axioms, | 
| 1301 | binarize) | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1302 | end | 
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1303 | |
| 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: diff
changeset | 1304 | end; |