src/HOL/Tools/refute.ML
author webertj
Fri, 26 Mar 2004 14:53:17 +0100
changeset 14488 863258b0cdca
parent 14460 04e787a4f17a
child 14604 1946097f7068
permissions -rw-r--r--
slightly different SAT solver interface
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     1
(*  Title:      HOL/Tools/refute.ML
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     2
    ID:         $Id$
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     3
    Author:     Tjark Weber
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     4
    Copyright   2003-2004
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     5
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
     6
Finite model generation for HOL formulae, using a SAT solver.
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     7
*)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     8
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     9
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    10
(* TODO: Change parameter handling so that only supported parameters can be  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    11
(*       set, and specify defaults here, rather than in HOL/Main.thy.        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    12
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    13
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    14
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    15
(* Declares the 'REFUTE' signature as well as a structure 'Refute'.          *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    16
(* Documentation is available in the Isabelle/Isar theory 'HOL/Refute.thy'.  *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    17
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    18
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    19
signature REFUTE =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    20
sig
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    21
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    22
	exception REFUTE of string * string
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    23
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    24
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    25
(* Model/interpretation related code (translation HOL -> boolean formula)    *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    26
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    27
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    28
	exception CANNOT_INTERPRET
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    29
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    30
	type interpretation
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    31
	type model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    32
	type arguments
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    33
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    34
	val add_interpreter : string -> (theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option) -> theory -> theory
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    35
	val add_printer     : string -> (theory -> model -> Term.term -> interpretation -> (int -> bool) -> string option) -> theory -> theory
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    36
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    37
	val interpret           : theory -> model -> arguments -> Term.term -> interpretation * model * arguments  (* exception CANNOT_INTERPRET *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    38
	val is_satisfying_model : model -> interpretation -> bool -> PropLogic.prop_formula
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    39
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    40
	val print       : theory -> model -> Term.term -> interpretation -> (int -> bool) -> string
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    41
	val print_model : theory -> model -> (int -> bool) -> string
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    42
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    43
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    44
(* Interface                                                                 *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    45
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    46
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    47
	val set_default_param  : (string * string) -> theory -> theory
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    48
	val get_default_param  : theory -> string -> string option
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    49
	val get_default_params : theory -> (string * string) list
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    50
	val actual_params      : theory -> (string * string) list -> (int * int * int * string)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    51
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    52
	val find_model : theory -> (int * int * int * string) -> Term.term -> bool -> unit
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    53
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    54
	val satisfy_term   : theory -> (string * string) list -> Term.term -> unit  (* tries to find a model for a formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    55
	val refute_term    : theory -> (string * string) list -> Term.term -> unit  (* tries to find a model that refutes a formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    56
	val refute_subgoal : theory -> (string * string) list -> Thm.thm -> int -> unit
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    57
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    58
	val setup : (theory -> theory) list
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    59
end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    60
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    61
structure Refute : REFUTE =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    62
struct
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    63
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    64
	open PropLogic;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    65
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    66
	(* We use 'REFUTE' only for internal error conditions that should    *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    67
	(* never occur in the first place (i.e. errors caused by bugs in our *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    68
	(* code).  Otherwise (e.g. to indicate invalid input data) we use    *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    69
	(* 'error'.                                                          *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    70
	exception REFUTE of string * string;  (* ("in function", "cause") *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    71
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    72
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    73
(* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL       *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    74
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    75
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    76
	exception CANNOT_INTERPRET;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    77
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    78
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    79
(* TREES                                                                     *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    80
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    81
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    82
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    83
(* tree: implements an arbitrarily (but finitely) branching tree as a list   *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    84
(*       of (lists of ...) elements                                          *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    85
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    86
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    87
	datatype 'a tree =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    88
		  Leaf of 'a
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    89
		| Node of ('a tree) list;
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    90
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    91
	(* ('a -> 'b) -> 'a tree -> 'b tree *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    92
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    93
	fun tree_map f tr =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    94
		case tr of
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    95
		  Leaf x  => Leaf (f x)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    96
		| Node xs => Node (map (tree_map f) xs);
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    97
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    98
	(* ('a * 'b -> 'a) -> 'a * ('b tree) -> 'a *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    99
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   100
	fun tree_foldl f =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   101
	let
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   102
		fun itl (e, Leaf x)  = f(e,x)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   103
		  | itl (e, Node xs) = foldl (tree_foldl f) (e,xs)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   104
	in
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   105
		itl
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   106
	end;
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   107
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   108
	(* 'a tree * 'b tree -> ('a * 'b) tree *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   109
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   110
	fun tree_pair (t1,t2) =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   111
		case t1 of
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   112
		  Leaf x =>
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   113
			(case t2 of
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   114
				  Leaf y => Leaf (x,y)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   115
				| Node _ => raise REFUTE ("tree_pair", "trees are of different height (second tree is higher)"))
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   116
		| Node xs =>
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   117
			(case t2 of
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   118
				  (* '~~' will raise an exception if the number of branches in both trees is different at the current node *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   119
				  Node ys => Node (map tree_pair (xs ~~ ys))
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   120
				| Leaf _  => raise REFUTE ("tree_pair", "trees are of different height (first tree is higher)"));
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   121
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   122
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   123
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   124
(* interpretation: a term's interpretation is given by a variable of type    *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   125
(*                 'interpretation'                                          *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   126
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   127
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   128
	type interpretation =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   129
		prop_formula list tree;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   130
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   131
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   132
(* model: a model specifies the size of types and the interpretation of      *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   133
(*        terms                                                              *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   134
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   135
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   136
	type model =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   137
		(Term.typ * int) list * (Term.term * interpretation) list;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   138
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   139
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   140
(* arguments: additional arguments required during interpretation of terms   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   141
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   142
	
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   143
	type arguments =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   144
		{next_idx: int, bounds: interpretation list, wellformed: prop_formula};
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   145
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   146
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   147
	structure RefuteDataArgs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   148
	struct
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   149
		val name = "HOL/refute";
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   150
		type T =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   151
			{interpreters: (string * (theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option)) list,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   152
			 printers: (string * (theory -> model -> Term.term -> interpretation -> (int -> bool) -> string option)) list,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   153
			 parameters: string Symtab.table};
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   154
		val empty = {interpreters = [], printers = [], parameters = Symtab.empty};
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   155
		val copy = I;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   156
		val prep_ext = I;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   157
		fun merge
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   158
			({interpreters = in1, printers = pr1, parameters = pa1},
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   159
			 {interpreters = in2, printers = pr2, parameters = pa2}) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   160
			{interpreters = rev (merge_alists (rev in1) (rev in2)),
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   161
			 printers = rev (merge_alists (rev pr1) (rev pr2)),
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   162
			 parameters = Symtab.merge (op=) (pa1, pa2)};
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   163
		fun print sg {interpreters, printers, parameters} =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   164
			Pretty.writeln (Pretty.chunks
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   165
				[Pretty.strs ("default parameters:" :: flat (map (fn (name,value) => [name, "=", value]) (Symtab.dest parameters))),
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   166
				 Pretty.strs ("interpreters:" :: map fst interpreters),
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   167
				 Pretty.strs ("printers:" :: map fst printers)]);
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   168
	end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   169
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   170
	structure RefuteData = TheoryDataFun(RefuteDataArgs);
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   171
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   172
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   173
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   174
(* interpret: tries to interpret the term 't' using a suitable interpreter;  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   175
(*            returns the interpretation and a (possibly extended) model     *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   176
(*            that keeps track of the interpretation of subterms             *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   177
(* Note: exception 'CANNOT_INTERPRETE' is raised if the term cannot be       *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   178
(*       interpreted by any interpreter                                      *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   179
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   180
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   181
	(* theory -> model -> arguments -> Term.term -> interpretation * model * arguments *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   182
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   183
	fun interpret thy model args t =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   184
		(case get_first (fn (_, f) => f thy model args t) (#interpreters (RefuteData.get thy)) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   185
		  None =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   186
			(std_output "\n"; warning ("Unable to interpret term:\n" ^ Sign.string_of_term (sign_of thy) t);
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   187
			raise CANNOT_INTERPRET)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   188
		| Some x => x);
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   189
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   190
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   191
(* print: tries to print the constant denoted by the term 't' using a        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   192
(*        suitable printer                                                   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   193
(* ------------------------------------------------------------------------- *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   194
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   195
	(* theory -> model -> Term.term -> interpretation -> (int -> bool) -> string *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   196
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   197
	fun print thy model t intr assignment =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   198
		(case get_first (fn (_, f) => f thy model t intr assignment) (#printers (RefuteData.get thy)) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   199
		  None => "<<no printer available>>"
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   200
		| Some s => s);
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   201
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   202
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   203
(* is_satisfying_model: returns a propositional formula that is true iff the *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   204
(*      given interpretation denotes 'satisfy', and the model meets certain  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   205
(*      well-formedness properties                                           *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   206
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   207
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   208
	(* model -> interpretation -> bool -> PropLogic.prop_formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   209
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   210
	fun is_satisfying_model model intr satisfy =
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   211
	let
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   212
		(* prop_formula list -> prop_formula *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   213
		fun allfalse []      = True
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   214
		  | allfalse (x::xs) = SAnd (SNot x, allfalse xs)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   215
		(* prop_formula list -> prop_formula *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   216
		fun exactly1true []      = False
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   217
		  | exactly1true (x::xs) = SOr (SAnd (x, allfalse xs), SAnd (SNot x, exactly1true xs))
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   218
	(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   219
	(* restrict_to_single_element: returns a propositional formula that is true  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   220
	(*      iff the interpretation denotes a single element of its corresponding *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   221
	(*      type, i.e. iff at each leaf, one and only one formula is true        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   222
	(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   223
		(* interpretation -> prop_formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   224
		fun restrict_to_single_element (Leaf [BoolVar i, Not (BoolVar j)]) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   225
			if (i=j) then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   226
				True  (* optimization for boolean variables *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   227
			else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   228
				raise REFUTE ("is_satisfying_model", "boolean variables in leaf have different indices")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   229
		  | restrict_to_single_element (Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   230
			exactly1true xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   231
		  | restrict_to_single_element (Node trees) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   232
			PropLogic.all (map restrict_to_single_element trees)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   233
	in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   234
		(* a term of type 'bool' is represented as a 2-element leaf, where  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   235
		(* the term is true iff the leaf's first element is true, and false *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   236
		(* iff the leaf's second element is true                            *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   237
		case intr of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   238
		  Leaf [fmT,fmF] =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   239
			(* well-formedness properties and formula 'fmT'/'fmF' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   240
			SAnd (PropLogic.all (map (restrict_to_single_element o snd) (snd model)), if satisfy then fmT else fmF)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   241
		| _ =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   242
			raise REFUTE ("is_satisfying_model", "interpretation does not denote a boolean value")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   243
	end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   244
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   245
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   246
(* print_model: turns the model into a string, using a fixed interpretation  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   247
(*              (given by an assignment for boolean variables) and suitable  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   248
(*              printers                                                     *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   249
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   250
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   251
	fun print_model thy model assignment =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   252
	let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   253
		val (typs, terms) = model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   254
	in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   255
		(if null typs then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   256
			"empty universe (no type variables in term)"
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   257
		else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   258
			"Size of types: " ^ commas (map (fn (T,i) => Sign.string_of_typ (sign_of thy) T ^ ": " ^ string_of_int i) typs))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   259
		^ "\n"
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   260
		^ (if null terms then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   261
			"empty interpretation (no free variables in term)"
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   262
		  else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   263
			space_implode "\n" (map
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   264
				(fn (t,intr) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   265
					Sign.string_of_term (sign_of thy) t ^ ": " ^ print thy model t intr assignment)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   266
				terms)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   267
		  )
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   268
		^ "\n"
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   269
	end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   270
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   271
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   272
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   273
(* PARAMETER MANAGEMENT                                                      *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   274
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   275
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   276
	(* string -> (theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option) -> theory -> theory *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   277
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   278
	fun add_interpreter name f thy =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   279
	let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   280
		val {interpreters, printers, parameters} = RefuteData.get thy
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   281
	in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   282
		case assoc (interpreters, name) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   283
		  None   => RefuteData.put {interpreters = (name, f) :: interpreters, printers = printers, parameters = parameters} thy
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   284
		| Some _ => error ("Interpreter " ^ name ^ " already declared")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   285
	end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   286
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   287
	(* string -> (theory -> model -> Term.term -> interpretation -> (int -> bool) -> string option) -> theory -> theory *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   288
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   289
	fun add_printer name f thy =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   290
	let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   291
		val {interpreters, printers, parameters} = RefuteData.get thy
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   292
	in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   293
		case assoc (printers, name) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   294
		  None   => RefuteData.put {interpreters = interpreters, printers = (name, f) :: printers, parameters = parameters} thy
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   295
		| Some _ => error ("Printer " ^ name ^ " already declared")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   296
	end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   297
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   298
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   299
(* set_default_param: stores the '(name, value)' pair in RefuteData's        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   300
(*                    parameter table                                        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   301
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   302
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   303
	(* (string * string) -> theory -> theory *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   304
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   305
	fun set_default_param (name, value) thy =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   306
	let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   307
		val {interpreters, printers, parameters} = RefuteData.get thy
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   308
	in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   309
		case Symtab.lookup (parameters, name) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   310
		  None   => RefuteData.put
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   311
			{interpreters = interpreters, printers = printers, parameters = Symtab.extend (parameters, [(name, value)])} thy
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   312
		| Some _ => RefuteData.put
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   313
			{interpreters = interpreters, printers = printers, parameters = Symtab.update ((name, value), parameters)} thy
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   314
	end;
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   315
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   316
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   317
(* get_default_param: retrieves the value associated with 'name' from        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   318
(*                    RefuteData's parameter table                           *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   319
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   320
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   321
	(* theory -> string -> string option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   322
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   323
	fun get_default_param thy name = Symtab.lookup ((#parameters o RefuteData.get) thy, name);
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   324
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   325
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   326
(* get_default_params: returns a list of all '(name, value)' pairs that are  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   327
(*                     stored in RefuteData's parameter table                *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   328
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   329
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   330
	(* theory -> (string * string) list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   331
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   332
	fun get_default_params thy = (Symtab.dest o #parameters o RefuteData.get) thy;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   333
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   334
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   335
(* actual_params: takes a (possibly empty) list 'params' of parameters that  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   336
(*      override the default parameters currently specified in 'thy', and    *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   337
(*      returns a tuple that can be passed to 'find_model'.                  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   338
(*                                                                           *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   339
(* The following parameters are supported (and required!):                   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   340
(*                                                                           *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   341
(* Name          Type    Description                                         *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   342
(*                                                                           *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   343
(* "minsize"     int     Only search for models with size at least           *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   344
(*                       'minsize'.                                          *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   345
(* "maxsize"     int     If >0, only search for models with size at most     *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   346
(*                       'maxsize'.                                          *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   347
(* "maxvars"     int     If >0, use at most 'maxvars' boolean variables      *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   348
(*                       when transforming the term into a propositional     *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   349
(*                       formula.                                            *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   350
(* "satsolver"   string  SAT solver to be used.                              *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   351
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   352
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   353
	(* theory -> (string * string) list -> (int * int * int * string) *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   354
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   355
	fun actual_params thy params =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   356
	let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   357
		(* (string * string) list * string -> int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   358
		fun read_int (parms, name) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   359
			case assoc_string (parms, name) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   360
			  Some s => (case Int.fromString s of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   361
				  SOME i => i
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   362
				| NONE   => error ("parameter " ^ quote name ^ " (value is " ^ quote s ^ ") must be an integer value"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   363
			| None   => error ("parameter " ^ quote name ^ " must be assigned a value")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   364
		(* (string * string) list * string -> string *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   365
		fun read_string (parms, name) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   366
			case assoc_string (parms, name) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   367
			  Some s => s
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   368
			| None   => error ("parameter " ^ quote name ^ " must be assigned a value")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   369
		(* (string * string) list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   370
		val allparams = params @ (get_default_params thy)  (* 'params' first, defaults last (to override) *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   371
		(* int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   372
		val minsize   = read_int (allparams, "minsize")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   373
		val maxsize   = read_int (allparams, "maxsize")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   374
		val maxvars   = read_int (allparams, "maxvars")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   375
		(* string *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   376
		val satsolver = read_string (allparams, "satsolver")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   377
	in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   378
		(minsize, maxsize, maxvars, satsolver)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   379
	end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   380
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   381
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   382
(* find_model: repeatedly calls 'invoke_propgen' with appropriate            *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   383
(*             parameters, applies the SAT solver, and (in case a model is   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   384
(*             found) displays the model to the user                         *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   385
(* thy      : the current theory                                             *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   386
(* minsize  : the minimal size of the model                                  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   387
(* maxsize  : the maximal size of the model                                  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   388
(* maxvars  : the maximal number of boolean variables to be used             *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   389
(* satsolver: the name of the SAT solver to be used                          *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   390
(* satisfy  : if 'True', search for a model that makes 't' true; otherwise   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   391
(*            search for a model that makes 't' false                        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   392
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   393
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   394
	(* theory ->  (int * int * int * string) -> Term.term -> bool -> unit *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   395
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   396
	fun find_model thy (minsize, maxsize, maxvars, satsolver) t satisfy =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   397
	let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   398
		(* Term.typ list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   399
		val tvars  = map (fn (i,s) => TVar(i,s)) (term_tvars t)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   400
		val tfrees = map (fn (x,s) => TFree(x,s)) (term_tfrees t)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   401
		(* int -> unit *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   402
		fun find_model_loop size =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   403
			(* 'maxsize=0' means "search for arbitrary large models" *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   404
			if size>maxsize andalso maxsize<>0 then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   405
				writeln ("Search terminated: maxsize=" ^ string_of_int maxsize ^ " exceeded.")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   406
			else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   407
			let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   408
				(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   409
				(* make_universes: given a list 'xs' of "types" and a universe size 'size',  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   410
				(*      this function returns all possible partitions of the universe into   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   411
				(*      the "types" in 'xs' such that no "type" is empty.  If 'size' is less *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   412
				(*      than 'length xs', the returned list of partitions is empty.          *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   413
				(*      Otherwise, if the list 'xs' is empty, then the returned list of      *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   414
				(*      partitions contains only the empty list, regardless of 'size'.       *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   415
				(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   416
				(* 'a list -> int -> ('a * int) list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   417
				fun make_universes xs size =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   418
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   419
					(* 'a list -> int -> int -> ('a * int) list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   420
					fun make_partitions_loop (x::xs) 0 total =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   421
						map (fn us => ((x,0)::us)) (make_partitions xs total)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   422
					  | make_partitions_loop (x::xs) first total =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   423
						(map (fn us => ((x,first)::us)) (make_partitions xs (total-first))) @ (make_partitions_loop (x::xs) (first-1) total)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   424
					  | make_partitions_loop _ _ _ =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   425
						raise REFUTE ("find_model", "empty list (in make_partitions_loop)")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   426
					and
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   427
					(* 'a list -> int -> ('a * int) list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   428
					make_partitions [x] size =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   429
						(* we must use all remaining elements on the type 'x', so there is only one partition *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   430
						[[(x,size)]]
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   431
					  | make_partitions (x::xs) 0 =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   432
						(* there are no elements left in the universe, so there is only one partition *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   433
						[map (fn t => (t,0)) (x::xs)]
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   434
					  | make_partitions (x::xs) size =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   435
						(* we assign either size, size-1, ..., 1 or 0 elements to 'x'; the remaining elements are partitioned recursively *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   436
						make_partitions_loop (x::xs) size size
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   437
					  | make_partitions _ _ =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   438
						raise REFUTE ("find_model", "empty list (in make_partitions)")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   439
					val len = length xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   440
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   441
					if size<len then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   442
						(* the universe isn't big enough to make every type non-empty *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   443
						[]
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   444
					else if xs=[] then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   445
						(* no types: return one universe, regardless of the size *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   446
						[[]]
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   447
					else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   448
						(* partition into possibly empty types, then add 1 element to each type *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   449
						map (fn us => map (fn (x,i) => (x,i+1)) us) (make_partitions xs (size-len))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   450
				end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   451
				(* ('a * int) list -> bool *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   452
				fun find_interpretation xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   453
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   454
					val init_model          = (xs, [])
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   455
					val init_args           = {next_idx = 1, bounds = [], wellformed = True}
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   456
					val (intr, model, args) = interpret thy init_model init_args t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   457
					val fm                  = SAnd (#wellformed args, is_satisfying_model model intr satisfy)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   458
					val usedvars            = maxidx fm
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   459
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   460
					(* 'maxvars=0' means "use as many variables as necessary" *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   461
					if usedvars>maxvars andalso maxvars<>0 then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   462
						(writeln ("\nSearch terminated: " ^ string_of_int usedvars ^ " boolean variables used (only " ^ string_of_int maxvars ^ " allowed).");
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   463
						true)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   464
					else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   465
					(
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   466
						std_output " invoking SAT solver...";
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   467
						case SatSolver.invoke_solver satsolver fm of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   468
						  None =>
14488
863258b0cdca slightly different SAT solver interface
webertj
parents: 14460
diff changeset
   469
							(std_output (" error: SAT solver " ^ quote satsolver ^ " not configured.\n");
863258b0cdca slightly different SAT solver interface
webertj
parents: 14460
diff changeset
   470
							true)
863258b0cdca slightly different SAT solver interface
webertj
parents: 14460
diff changeset
   471
						| Some None =>
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   472
							(std_output " no model found.\n";
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   473
							false)
14488
863258b0cdca slightly different SAT solver interface
webertj
parents: 14460
diff changeset
   474
						| Some (Some assignment) =>
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   475
							(writeln ("\nModel found:\n" ^ print_model thy model assignment);
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   476
							true)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   477
					)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   478
				end handle CANNOT_INTERPRET => true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   479
					(* TODO: change this to false once the ability to interpret terms depends on the universe *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   480
			in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   481
				case make_universes (tvars@tfrees) size of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   482
				  [] =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   483
					(writeln ("Searching for a model of size " ^ string_of_int size ^ ": cannot make every type non-empty (model is too small).");
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   484
					find_model_loop (size+1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   485
				| [[]] =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   486
					(std_output ("Searching for a model of size " ^ string_of_int size ^ ", translating term...");
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   487
					if find_interpretation [] then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   488
						()
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   489
					else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   490
						writeln ("Search terminated: no type variables in term."))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   491
				| us =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   492
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   493
						fun loop []      =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   494
							find_model_loop (size+1)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   495
						  | loop (u::us) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   496
							(std_output ("Searching for a model of size " ^ string_of_int size ^ ", translating term...");
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   497
							if find_interpretation u then () else loop us)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   498
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   499
						loop us
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   500
					end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   501
			end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   502
	in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   503
		writeln ("Trying to find a model that "
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   504
			^ (if satisfy then "satisfies" else "refutes")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   505
			^ ": " ^ (Sign.string_of_term (sign_of thy) t));
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   506
		if minsize<1 then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   507
			writeln "\"minsize\" is less than 1; starting search with size 1."
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   508
		else ();
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   509
		find_model_loop (Int.max (minsize,1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   510
	end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   511
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   512
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   513
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   514
(* INTERFACE, PART 2: FINDING A MODEL                                        *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   515
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   516
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   517
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   518
(* satisfy_term: calls 'find_model' to find a model that satisfies 't'       *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   519
(* params      : list of '(name, value)' pairs used to override default      *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   520
(*               parameters                                                  *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   521
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   522
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   523
	(* theory -> (string * string) list -> Term.term -> unit *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   524
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   525
	fun satisfy_term thy params t =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   526
		find_model thy (actual_params thy params) t true;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   527
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   528
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   529
(* refute_term: calls 'find_model' to find a model that refutes 't'          *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   530
(* params     : list of '(name, value)' pairs used to override default       *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   531
(*              parameters                                                   *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   532
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   533
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   534
	(* theory -> (string * string) list -> Term.term -> unit *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   535
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   536
	fun refute_term thy params t =
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   537
	let
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   538
		(* disallow schematic type variables, since we cannot properly negate terms containing them *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   539
		val _ = assert (null (term_tvars t)) "Term to be refuted contains schematic type variables"
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   540
		(* existential closure over schematic variables *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   541
		(* (Term.indexname * Term.typ) list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   542
		val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   543
		(* Term.term *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   544
		val ex_closure = foldl
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   545
			(fn (t', ((x,i),T)) => (HOLogic.exists_const T) $ Abs (x, T, abstract_over (Var((x,i),T), t')))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   546
			(t, vars)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   547
		(* If 't' is of type 'propT' (rather than 'boolT'), applying  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   548
		(* 'HOLogic.exists_const' is not type-correct.  However, this *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   549
		(* isn't really a problem as long as 'find_model' still       *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   550
		(* interprets the resulting term correctly, without checking  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   551
		(* its type.                                                  *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   552
	in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   553
		find_model thy (actual_params thy params) ex_closure false
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   554
	end;
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   555
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   556
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   557
(* refute_subgoal: calls 'refute_term' on a specific subgoal                 *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   558
(* params        : list of '(name, value)' pairs used to override default    *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   559
(*                 parameters                                                *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   560
(* subgoal       : 0-based index specifying the subgoal number               *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   561
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   562
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   563
	(* theory -> (string * string) list -> Thm.thm -> int -> unit *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   564
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   565
	fun refute_subgoal thy params thm subgoal =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   566
		refute_term thy params (nth_elem (subgoal, prems_of thm));
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   567
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   568
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   569
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   570
(* INTERPRETERS                                                              *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   571
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   572
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   573
	(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   574
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   575
	fun var_typevar_interpreter thy model args t =
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   576
	let
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   577
		fun is_var (Free _) = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   578
		  | is_var (Var _)  = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   579
		  | is_var _        = false
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   580
		fun typeof (Free (_,T)) = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   581
		  | typeof (Var (_,T))  = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   582
		  | typeof _            = raise REFUTE ("var_typevar_interpreter", "term is not a variable")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   583
		fun is_typevar (TFree _) = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   584
		  | is_typevar (TVar _)  = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   585
		  | is_typevar _         = false
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   586
		val (sizes, intrs) = model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   587
	in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   588
		if is_var t andalso is_typevar (typeof t) then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   589
			(case assoc (intrs, t) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   590
			  Some intr => Some (intr, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   591
			| None      =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   592
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   593
					val size = (the o assoc) (sizes, typeof t)  (* the model MUST specify a size for type variables *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   594
					val idx  = #next_idx args
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   595
					val intr = (if size=2 then Leaf [BoolVar idx, Not (BoolVar idx)] else Leaf (map BoolVar (idx upto (idx+size-1))))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   596
					val next = (if size=2 then idx+1 else idx+size)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   597
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   598
					(* extend the model, increase 'next_idx' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   599
					Some (intr, (sizes, (t, intr)::intrs), {next_idx = next, bounds = #bounds args, wellformed = #wellformed args})
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   600
				end)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   601
		else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   602
			None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   603
	end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   604
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   605
	(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   606
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   607
	fun var_funtype_interpreter thy model args t =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   608
	let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   609
		fun is_var (Free _) = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   610
		  | is_var (Var _)  = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   611
		  | is_var _        = false
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   612
		fun typeof (Free (_,T)) = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   613
		  | typeof (Var (_,T))  = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   614
		  | typeof _            = raise REFUTE ("var_funtype_interpreter", "term is not a variable")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   615
		fun stringof (Free (x,_))     = x
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   616
		  | stringof (Var ((x,_), _)) = x
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   617
		  | stringof _                = raise REFUTE ("var_funtype_interpreter", "term is not a variable")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   618
		fun is_funtype (Type ("fun", [_,_])) = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   619
		  | is_funtype _                     = false
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   620
		val (sizes, intrs) = model
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   621
	in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   622
		if is_var t andalso is_funtype (typeof t) then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   623
			(case assoc (intrs, t) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   624
			  Some intr => Some (intr, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   625
			| None      =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   626
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   627
					val T1 = domain_type (typeof t)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   628
					val T2 = range_type (typeof t)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   629
					(* we create 'size_of_interpretation (interpret (... T1))' different copies *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   630
					(* of the tree for 'T2', which are then combined into a single new tree     *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   631
					val (i1, _, _) = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free (stringof t, T1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   632
					(* power(a,b) computes a^b, for a>=0, b>=0 *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   633
					(* int * int -> int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   634
					fun power (a,0) = 1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   635
					  | power (a,1) = a
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   636
					  | power (a,b) = let val ab = power(a,b div 2) in ab * ab * power(a,b mod 2) end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   637
					fun size_of_interpretation (Leaf xs) = length xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   638
					  | size_of_interpretation (Node xs) = power (size_of_interpretation (hd xs), length xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   639
					val size = size_of_interpretation i1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   640
					(* make fresh copies, with different variable indices *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   641
					(* int -> int -> (int * interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   642
					fun make_copies idx 0 =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   643
						(idx, [])
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   644
					  | make_copies idx n =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   645
						let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   646
							val (copy, _, args) = interpret thy (sizes, []) {next_idx = idx, bounds = [], wellformed=True} (Free (stringof t, T2))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   647
							val (next, copies)  = make_copies (#next_idx args) (n-1)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   648
						in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   649
							(next, copy :: copies)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   650
						end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   651
					val (idx, copies) = make_copies (#next_idx args) (size_of_interpretation i1)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   652
					(* combine copies into a single tree *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   653
					val intr = Node copies
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   654
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   655
					(* extend the model, increase 'next_idx' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   656
					Some (intr, (sizes, (t, intr)::intrs), {next_idx = idx, bounds = #bounds args, wellformed = #wellformed args})
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   657
				end)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   658
		else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   659
			None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   660
	end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   661
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   662
	(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   663
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   664
	fun var_settype_interpreter thy model args t =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   665
		let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   666
			val (sizes, intrs) = model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   667
		in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   668
			case t of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   669
			  Free (x, Type ("set", [T])) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   670
				(case assoc (intrs, t) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   671
				  Some intr => Some (intr, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   672
				| None      =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   673
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   674
						val (intr, _, args') = interpret thy model args (Free (x, Type ("fun", [T, Type ("bool", [])])))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   675
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   676
						Some (intr, (sizes, (t, intr)::intrs), args')
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   677
					end)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   678
			| Var ((x,i), Type ("set", [T])) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   679
				(case assoc (intrs, t) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   680
				  Some intr => Some (intr, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   681
				| None      =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   682
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   683
						val (intr, _, args') = interpret thy model args (Var ((x,i), Type ("fun", [T, Type ("bool", [])])))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   684
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   685
						Some (intr, (sizes, (t, intr)::intrs), args')
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   686
					end)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   687
			| _ => None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   688
		end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   689
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   690
	(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   691
14460
04e787a4f17a SML/NJ compatibility fixes
webertj
parents: 14456
diff changeset
   692
	fun boundvar_interpreter thy model args (Bound i) = Some (nth_elem (i, #bounds (args:arguments)), model, args)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   693
	  | boundvar_interpreter thy model args _         = None;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   694
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   695
	(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   696
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   697
	fun abstraction_interpreter thy model args (Abs (x, T, t)) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   698
		let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   699
			val (sizes, intrs) = model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   700
			(* create all constants of type T *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   701
			val (i, _, _) = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free (x, T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   702
			(* returns a list with all unit vectors of length n *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   703
			(* int -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   704
			fun unit_vectors n =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   705
			let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   706
				(* returns the k-th unit vector of length n *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   707
				(* int * int -> interpretation *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   708
				fun unit_vector (k,n) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   709
					Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   710
				(* int -> interpretation list -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   711
				fun unit_vectors_acc k vs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   712
					if k>n then [] else (unit_vector (k,n))::(unit_vectors_acc (k+1) vs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   713
			in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   714
				unit_vectors_acc 1 []
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   715
			end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   716
			(* concatenates 'x' with every list in 'xss', returning a new list of lists *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   717
			(* 'a -> 'a list list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   718
			fun cons_list x xss =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   719
				map (fn xs => x::xs) xss
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   720
			(* returns a list of lists, each one consisting of n (possibly identical) elements from 'xs' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   721
			(* int -> 'a list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   722
			fun pick_all 1 xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   723
				map (fn x => [x]) xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   724
			  | pick_all n xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   725
				let val rec_pick = pick_all (n-1) xs in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   726
					foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   727
				end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   728
			(* interpretation -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   729
			fun make_constants (Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   730
				unit_vectors (length xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   731
			  | make_constants (Node xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   732
				map (fn xs' => Node xs') (pick_all (length xs) (make_constants (hd xs)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   733
			(* interpret the body 't' separately for each constant *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   734
			val ((model', args'), bodies) = foldl_map
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   735
				(fn ((m,a), c) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   736
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   737
						(* add 'c' to 'bounds' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   738
						val (i',m',a') = interpret thy m {next_idx = #next_idx a, bounds = c::(#bounds a), wellformed = #wellformed a} t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   739
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   740
						(* keep the new model m' and 'next_idx', but use old 'bounds' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   741
						((m',{next_idx = #next_idx a', bounds = #bounds a, wellformed = #wellformed a'}), i')
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   742
					end)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   743
				((model,args), make_constants i)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   744
		in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   745
			Some (Node bodies, model', args')
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   746
		end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   747
	  | abstraction_interpreter thy model args _               = None;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   748
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   749
	(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   750
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   751
	fun apply_interpreter thy model args (t1 $ t2) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   752
		let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   753
			(* auxiliary functions *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   754
			(* interpretation * interpretation -> interpretation *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   755
			fun interpretation_disjunction (tr1,tr2) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   756
				tree_map (fn (xs,ys) => map (fn (x,y) => SOr(x,y)) (xs ~~ ys)) (tree_pair (tr1,tr2))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   757
			(* prop_formula * interpretation -> interpretation *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   758
			fun prop_formula_times_interpretation (fm,tr) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   759
				tree_map (map (fn x => SAnd (fm,x))) tr
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   760
			(* prop_formula list * interpretation list -> interpretation *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   761
			fun prop_formula_list_dot_product_interpretation_list ([fm],[tr]) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   762
				prop_formula_times_interpretation (fm,tr)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   763
			  | prop_formula_list_dot_product_interpretation_list (fm::fms,tr::trees) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   764
				interpretation_disjunction (prop_formula_times_interpretation (fm,tr), prop_formula_list_dot_product_interpretation_list (fms,trees))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   765
			  | prop_formula_list_dot_product_interpretation_list (_,_) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   766
				raise REFUTE ("apply_interpreter", "empty list (in dot product)")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   767
			(* concatenates 'x' with every list in 'xss', returning a new list of lists *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   768
			(* 'a -> 'a list list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   769
			fun cons_list x xss =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   770
				map (fn xs => x::xs) xss
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   771
			(* returns a list of lists, each one consisting of one element from each element of 'xss' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   772
			(* 'a list list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   773
			fun pick_all [xs] =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   774
				map (fn x => [x]) xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   775
			  | pick_all (xs::xss) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   776
				let val rec_pick = pick_all xss in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   777
					foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   778
				end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   779
			  | pick_all _ =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   780
				raise REFUTE ("apply_interpreter", "empty list (in pick_all)")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   781
			(* interpretation -> prop_formula list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   782
			fun interpretation_to_prop_formula_list (Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   783
				xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   784
			  | interpretation_to_prop_formula_list (Node trees) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   785
				map PropLogic.all (pick_all (map interpretation_to_prop_formula_list trees))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   786
			(* interpretation * interpretation -> interpretation *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   787
			fun interpretation_apply (tr1,tr2) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   788
				(case tr1 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   789
				  Leaf _ =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   790
					raise REFUTE ("apply_interpreter", "first interpretation is a leaf")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   791
				| Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   792
					prop_formula_list_dot_product_interpretation_list (interpretation_to_prop_formula_list tr2, xs))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   793
			(* interpret 't1' and 't2' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   794
			val (intr1, model1, args1) = interpret thy model args t1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   795
			val (intr2, model2, args2) = interpret thy model1 args1 t2
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   796
		in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   797
			Some (interpretation_apply (intr1,intr2), model2, args2)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   798
		end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   799
	  | apply_interpreter thy model args _         = None;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   800
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   801
	(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   802
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   803
	fun const_unfold_interpreter thy model args t =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   804
		(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   805
		(* We unfold constants for which a defining equation is given as an axiom.   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   806
		(* A polymorphic axiom's type variables are instantiated.  Eta-expansion is  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   807
		(* performed only if necessary; arguments in the axiom that are present as   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   808
		(* actual arguments in 't' are simply substituted.  If more actual than      *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   809
		(* formal arguments are present, the constant is *not* unfolded, so that     *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   810
		(* other interpreters (that may just not have looked into the term far       *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   811
		(* enough yet) may be applied first (after 'apply_interpreter' gets rid of   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   812
		(* one argument).                                                            *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   813
		(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   814
		let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   815
			val (head, actuals) = strip_comb t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   816
			val actuals_count   = length actuals
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   817
		in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   818
			case head of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   819
			  Const (s,T) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   820
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   821
					(* (string * Term.term) list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   822
					val axioms = flat (map (Symtab.dest o #axioms o Theory.rep_theory) (thy :: Theory.ancestors_of thy))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   823
					(* Term.term * Term.term list * Term.term list -> Term.term *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   824
					(* term, formal arguments, actual arguments *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   825
					fun normalize (t, [],    [])    = t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   826
					  | normalize (t, [],    y::ys) = raise REFUTE ("const_unfold_interpreter", "more actual than formal parameters")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   827
					  | normalize (t, x::xs, [])    = normalize (lambda x t, xs, [])                (* eta-expansion *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   828
					  | normalize (t, x::xs, y::ys) = normalize (betapply (lambda x t, y), xs, ys)  (* substitution *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   829
					(* string -> Term.typ -> (string * Term.term) list -> Term.term option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   830
					fun get_defn s T [] =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   831
						None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   832
					  | get_defn s T ((_,ax)::axms) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   833
						(let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   834
							val (lhs, rhs)   = Logic.dest_equals ax  (* equations only *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   835
							val (c, formals) = strip_comb lhs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   836
							val (s', T')     = dest_Const c
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   837
						in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   838
							if (s=s') andalso (actuals_count <= length formals) then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   839
								let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   840
									val varT'      = Type.varifyT T'  (* for polymorphic definitions *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   841
									val typeSubs   = Type.typ_match (Sign.tsig_of (sign_of thy)) (Vartab.empty, (varT', T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   842
									val varRhs     = map_term_types Type.varifyT rhs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   843
									val varFormals = map (map_term_types Type.varifyT) formals
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   844
									val rhs'       = normalize (varRhs, varFormals, actuals)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   845
									val unvarRhs   = map_term_types
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   846
										(map_type_tvar
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   847
											(fn (v,_) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   848
												case Vartab.lookup (typeSubs, v) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   849
												  None =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   850
													raise REFUTE ("const_unfold_interpreter", "schematic type variable " ^ (fst v) ^ "not instantiated (in definition of " ^ quote s ^ ")")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   851
												| Some typ =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   852
													typ))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   853
										rhs'
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   854
								in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   855
									Some unvarRhs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   856
								end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   857
							else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   858
								get_defn s T axms
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   859
						end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   860
						handle TERM _          => get_defn s T axms
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   861
						     | Type.TYPE_MATCH => get_defn s T axms)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   862
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   863
					case get_defn s T axioms of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   864
					  Some t' => Some (interpret thy model args t')
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   865
					| None    => None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   866
				end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   867
			| _ => None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   868
		end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   869
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   870
	(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   871
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   872
	fun Pure_interpreter thy model args t =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   873
		let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   874
			fun toTrue (Leaf [fm,_]) = fm
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   875
			  | toTrue _             = raise REFUTE ("Pure_interpreter", "interpretation does not denote a boolean value")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   876
			fun toFalse (Leaf [_,fm]) = fm
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   877
			  | toFalse _             = raise REFUTE ("Pure_interpreter", "interpretation does not denote a boolean value")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   878
		in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   879
			case t of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   880
			  (*Const ("Goal", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   881
				Some (interpret thy model args t1) |*)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   882
			  Const ("all", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   883
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   884
					val (i,m,a) = (interpret thy model args t1)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   885
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   886
					case i of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   887
					  Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   888
						let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   889
							val fmTrue  = PropLogic.all (map toTrue xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   890
							val fmFalse = PropLogic.exists (map toFalse xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   891
						in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   892
							Some (Leaf [fmTrue, fmFalse], m, a)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   893
						end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   894
					| _ =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   895
						raise REFUTE ("Pure_interpreter", "\"all\" is not followed by a function")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   896
				end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   897
			| Const ("==", _) $ t1 $ t2 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   898
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   899
					val (i1,m1,a1) = interpret thy model args t1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   900
					val (i2,m2,a2) = interpret thy m1 a1 t2
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   901
					(* interpretation * interpretation -> prop_formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   902
					fun interpret_equal (tr1,tr2) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   903
						(case tr1 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   904
						  Leaf x =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   905
							(case tr2 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   906
							  Leaf y => PropLogic.dot_product (x,y)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   907
							| _      => raise REFUTE ("Pure_interpreter", "\"==\": second tree is higher"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   908
						| Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   909
							(case tr2 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   910
							  Leaf _  => raise REFUTE ("Pure_interpreter", "\"==\": first tree is higher")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   911
							(* extensionality: two functions are equal iff they are equal for every element *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   912
							| Node ys => PropLogic.all (map interpret_equal (xs ~~ ys))))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   913
					(* interpretation * interpretation -> prop_formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   914
					fun interpret_unequal (tr1,tr2) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   915
						(case tr1 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   916
						  Leaf x =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   917
							(case tr2 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   918
							  Leaf y =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   919
								let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   920
									fun unequal [] ([],_) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   921
										False
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   922
									  | unequal (x::xs) (y::ys1, ys2) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   923
										SOr (SAnd (x, PropLogic.exists (ys1@ys2)), unequal xs (ys1, y::ys2))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   924
									  | unequal _ _ =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   925
										raise REFUTE ("Pure_interpreter", "\"==\": leafs have different width")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   926
								in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   927
									unequal x (y,[])
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   928
								end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   929
							| _      => raise REFUTE ("Pure_interpreter", "\"==\": second tree is higher"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   930
						| Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   931
							(case tr2 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   932
							  Leaf _  => raise REFUTE ("Pure_interpreter", "\"==\": first tree is higher")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   933
							(* extensionality: two functions are unequal iff there exist unequal elements *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   934
							| Node ys => PropLogic.exists (map interpret_unequal (xs ~~ ys))))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   935
					val fmTrue  = interpret_equal (i1,i2)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   936
					val fmFalse = interpret_unequal (i1,i2)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   937
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   938
					Some (Leaf [fmTrue, fmFalse], m2, a2)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   939
				end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   940
			| Const ("==>", _) $ t1 $ t2 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   941
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   942
					val (i1,m1,a1) = interpret thy model args t1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   943
					val (i2,m2,a2) = interpret thy m1 a1 t2
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   944
					val fmTrue     = SOr (toFalse i1, toTrue i2)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   945
					val fmFalse    = SAnd (toTrue i1, toFalse i2)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   946
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   947
					Some (Leaf [fmTrue, fmFalse], m2, a2)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   948
				end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   949
			| _ => None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   950
		end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   951
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   952
	(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   953
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   954
	fun HOLogic_interpreter thy model args t =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   955
	(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   956
	(* Providing interpretations directly is more efficient than unfolding the   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   957
	(* logical constants; however, we need versions for constants with arguments *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   958
	(* (to avoid unfolding) as well as versions for constants without arguments  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   959
	(* (since in HOL, logical constants can themselves be arguments)             *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   960
	(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   961
	let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   962
		fun eta_expand t i =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   963
			let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   964
				val Ts = binder_types (fastype_of t)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   965
			in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   966
				foldr (fn (T, t) => Abs ("<eta_expand>", T, t))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   967
					(take (i, Ts), list_comb (t, map Bound (i-1 downto 0)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   968
			end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   969
		val TT = Leaf [True, False]
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   970
		val FF = Leaf [False, True]
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   971
		fun toTrue (Leaf [fm,_]) = fm
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   972
		  | toTrue _             = raise REFUTE ("HOLogic_interpreter", "interpretation does not denote a boolean value")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   973
		fun toFalse (Leaf [_,fm]) = fm
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   974
		  | toFalse _             = raise REFUTE ("HOLogic_interpreter", "interpretation does not denote a boolean value")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   975
	in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   976
		case t of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   977
		  Const ("Trueprop", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   978
			Some (interpret thy model args t1)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   979
		| Const ("Trueprop", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   980
			Some (Node [TT, FF], model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   981
		| Const ("Not", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   982
			apply_interpreter thy model args t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   983
		| Const ("Not", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   984
			Some (Node [FF, TT], model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   985
		| Const ("True", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   986
			Some (TT, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   987
		| Const ("False", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   988
			Some (FF, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   989
		| Const ("arbitrary", T) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   990
			(* treat 'arbitrary' just like a free variable of the same type *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   991
			(case assoc (snd model, t) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   992
			  Some intr =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   993
				Some (intr, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   994
			| None =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   995
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   996
					val (sizes, intrs) = model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   997
					val (intr, _, a)   = interpret thy (sizes, []) args (Free ("<arbitrary>", T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   998
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   999
					Some (intr, (sizes, (t,intr)::intrs), a)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1000
				end)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1001
		| Const ("The", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1002
			apply_interpreter thy model args t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1003
		| Const ("The", T as Type ("fun", [_, T'])) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1004
			(* treat 'The' like a free variable, but with a fixed interpretation for boolean *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1005
			(* functions that map exactly one constant of type T' to True                    *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1006
			(case assoc (snd model, t) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1007
				Some intr =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1008
					Some (intr, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1009
			| None =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1010
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1011
					val (sizes, intrs) = model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1012
					val (intr, _, a)   = interpret thy (sizes, []) args (Free ("<The>", T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1013
					(* create all constants of type T' => bool, ... *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1014
					val (intrFun, _, _) = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free ("<The>", Type ("fun", [T', Type ("bool",[])])))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1015
					(* ... and all constants of type T' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1016
					val (intrT', _, _) = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free ("<The>", T'))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1017
					(* returns a list with all unit vectors of length n *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1018
					(* int -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1019
					fun unit_vectors n =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1020
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1021
						(* returns the k-th unit vector of length n *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1022
						(* int * int -> interpretation *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1023
						fun unit_vector (k,n) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1024
							Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1025
						(* int -> interpretation list -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1026
						fun unit_vectors_acc k vs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1027
							if k>n then [] else (unit_vector (k,n))::(unit_vectors_acc (k+1) vs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1028
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1029
						unit_vectors_acc 1 []
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1030
					end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1031
					(* concatenates 'x' with every list in 'xss', returning a new list of lists *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1032
					(* 'a -> 'a list list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1033
					fun cons_list x xss =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1034
						map (fn xs => x::xs) xss
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1035
					(* returns a list of lists, each one consisting of n (possibly identical) elements from 'xs' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1036
					(* int -> 'a list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1037
					fun pick_all 1 xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1038
						map (fn x => [x]) xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1039
					  | pick_all n xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1040
						let val rec_pick = pick_all (n-1) xs in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1041
							foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1042
						end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1043
					(* interpretation -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1044
					fun make_constants (Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1045
						unit_vectors (length xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1046
					  | make_constants (Node xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1047
						map (fn xs' => Node xs') (pick_all (length xs) (make_constants (hd xs)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1048
					val constantsFun = make_constants intrFun
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1049
					val constantsT'  = make_constants intrT'
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1050
					(* interpretation -> interpretation list -> interpretation option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1051
					fun the_val (Node xs) cs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1052
						let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1053
							val TrueCount = foldl (fn (n,x) => if toTrue x = True then n+1 else n) (0,xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1054
							fun findThe (x::xs) (c::cs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1055
								if toTrue x = True then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1056
									c
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1057
								else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1058
									findThe xs cs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1059
							  | findThe _ _ =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1060
								raise REFUTE ("HOLogic_interpreter", "\"The\": not found")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1061
						in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1062
							if TrueCount=1 then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1063
								Some (findThe xs cs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1064
							else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1065
								None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1066
						end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1067
					  | the_val _ _ =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1068
						raise REFUTE ("HOLogic_interpreter", "\"The\": function interpreted as a leaf")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1069
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1070
					case intr of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1071
					  Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1072
						let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1073
							(* replace interpretation 'x' by the interpretation determined by 'f' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1074
							val intrThe = Node (map (fn (x,f) => if_none (the_val f constantsT') x) (xs ~~ constantsFun))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1075
						in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1076
							Some (intrThe, (sizes, (t,intrThe)::intrs), a)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1077
						end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1078
					| _ =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1079
						raise REFUTE ("HOLogic_interpreter", "\"The\": interpretation is a leaf")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1080
				end)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1081
		| Const ("Hilbert_Choice.Eps", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1082
			apply_interpreter thy model args t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1083
		| Const ("Hilbert_Choice.Eps", T as Type ("fun", [_, T'])) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1084
			(* treat 'The' like a free variable, but with a fixed interpretation for boolean *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1085
			(* functions that map exactly one constant of type T' to True                    *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1086
			(case assoc (snd model, t) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1087
				Some intr =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1088
					Some (intr, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1089
			| None =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1090
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1091
					val (sizes, intrs) = model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1092
					val (intr, _, a)   = interpret thy (sizes, []) args (Free ("<Eps>", T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1093
					(* create all constants of type T' => bool, ... *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1094
					val (intrFun, _, _) = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free ("<Eps>", Type ("fun", [T', Type ("bool",[])])))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1095
					(* ... and all constants of type T' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1096
					val (intrT', _, _) = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free ("<Eps>", T'))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1097
					(* returns a list with all unit vectors of length n *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1098
					(* int -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1099
					fun unit_vectors n =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1100
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1101
						(* returns the k-th unit vector of length n *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1102
						(* int * int -> interpretation *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1103
						fun unit_vector (k,n) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1104
							Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1105
						(* int -> interpretation list -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1106
						fun unit_vectors_acc k vs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1107
							if k>n then [] else (unit_vector (k,n))::(unit_vectors_acc (k+1) vs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1108
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1109
						unit_vectors_acc 1 []
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1110
					end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1111
					(* concatenates 'x' with every list in 'xss', returning a new list of lists *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1112
					(* 'a -> 'a list list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1113
					fun cons_list x xss =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1114
						map (fn xs => x::xs) xss
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1115
					(* returns a list of lists, each one consisting of n (possibly identical) elements from 'xs' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1116
					(* int -> 'a list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1117
					fun pick_all 1 xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1118
						map (fn x => [x]) xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1119
					  | pick_all n xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1120
						let val rec_pick = pick_all (n-1) xs in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1121
							foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1122
						end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1123
					(* interpretation -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1124
					fun make_constants (Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1125
						unit_vectors (length xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1126
					  | make_constants (Node xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1127
						map (fn xs' => Node xs') (pick_all (length xs) (make_constants (hd xs)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1128
					val constantsFun = make_constants intrFun
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1129
					val constantsT'  = make_constants intrT'
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1130
					(* interpretation -> interpretation list -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1131
					fun true_values (Node xs) cs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1132
						mapfilter (fn (x,c) => if toTrue x = True then Some c else None) (xs~~cs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1133
					  | true_values _ _ =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1134
						raise REFUTE ("HOLogic_interpreter", "\"Eps\": function interpreted as a leaf")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1135
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1136
					case intr of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1137
					  Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1138
						let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1139
							val (wf, intrsEps) = foldl_map (fn (w,(x,f)) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1140
								case true_values f constantsT' of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1141
								  []  => (w, x)  (* no value mapped to true -> no restriction *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1142
								| [c] => (w, c)  (* one value mapped to true -> that's the one *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1143
								| cs  =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1144
									let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1145
										(* interpretation * interpretation -> prop_formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1146
										fun interpret_equal (tr1,tr2) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1147
											(case tr1 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1148
											  Leaf x =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1149
												(case tr2 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1150
												  Leaf y => PropLogic.dot_product (x,y)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1151
												| _      => raise REFUTE ("HOLogic_interpreter", "\"Eps\": second tree is higher"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1152
												| Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1153
												(case tr2 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1154
												  Leaf _  => raise REFUTE ("HOLogic_interpreter", "\"Eps\": first tree is higher")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1155
												(* extensionality: two functions are equal iff they are equal for every element *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1156
												| Node ys => PropLogic.all (map interpret_equal (xs ~~ ys))))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1157
									in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1158
										(SAnd (w, PropLogic.exists (map (fn c => interpret_equal (x,c)) cs)), x)  (* impose restrictions on x *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1159
									end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1160
								)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1161
								(#wellformed a, xs ~~ constantsFun)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1162
							val intrEps = Node intrsEps
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1163
						in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1164
							Some (intrEps, (sizes, (t,intrEps)::intrs), {next_idx = #next_idx a, bounds = #bounds a, wellformed = wf})
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1165
						end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1166
					| _ =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1167
						raise REFUTE ("HOLogic_interpreter", "\"Eps\": interpretation is a leaf")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1168
				end)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1169
		| Const ("All", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1170
			let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1171
				val (i,m,a) = interpret thy model args t1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1172
			in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1173
				case i of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1174
				  Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1175
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1176
						val fmTrue  = PropLogic.all (map toTrue xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1177
						val fmFalse = PropLogic.exists (map toFalse xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1178
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1179
						Some (Leaf [fmTrue, fmFalse], m, a)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1180
					end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1181
				| _ =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1182
					raise REFUTE ("HOLogic_interpreter", "\"All\" is not followed by a function")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1183
			end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1184
		| Const ("All", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1185
			Some (interpret thy model args (eta_expand t 1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1186
		| Const ("Ex", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1187
			let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1188
				val (i,m,a) = interpret thy model args t1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1189
			in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1190
				case i of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1191
				  Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1192
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1193
						val fmTrue  = PropLogic.exists (map toTrue xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1194
						val fmFalse = PropLogic.all (map toFalse xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1195
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1196
						Some (Leaf [fmTrue, fmFalse], m, a)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1197
					end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1198
				| _ =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1199
					raise REFUTE ("HOLogic_interpreter", "\"Ex\" is not followed by a function")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1200
			end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1201
		| Const ("Ex", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1202
			Some (interpret thy model args (eta_expand t 1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1203
		| Const ("Ex1", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1204
			let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1205
				val (i,m,a) = interpret thy model args t1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1206
			in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1207
				case i of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1208
				  Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1209
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1210
						(* interpretation list -> prop_formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1211
						fun allfalse []      = True
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1212
						  | allfalse (x::xs) = SAnd (toFalse x, allfalse xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1213
						(* interpretation list -> prop_formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1214
						fun exactly1true []      = False
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1215
						  | exactly1true (x::xs) = SOr (SAnd (toTrue x, allfalse xs), SAnd (toFalse x, exactly1true xs))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1216
						(* interpretation list -> prop_formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1217
						fun atleast2true []      = False
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1218
						  | atleast2true (x::xs) = SOr (SAnd (toTrue x, PropLogic.exists (map toTrue xs)), atleast2true xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1219
						val fmTrue  = exactly1true xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1220
						val fmFalse = SOr (allfalse xs, atleast2true xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1221
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1222
						Some (Leaf [fmTrue, fmFalse], m, a)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1223
					end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1224
				| _ =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1225
					raise REFUTE ("HOLogic_interpreter", "\"Ex1\" is not followed by a function")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1226
			end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1227
		| Const ("Ex1", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1228
			Some (interpret thy model args (eta_expand t 1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1229
		| Const ("op =", _) $ t1 $ t2 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1230
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1231
					val (i1,m1,a1) = interpret thy model args t1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1232
					val (i2,m2,a2) = interpret thy m1 a1 t2
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1233
					(* interpretation * interpretation -> prop_formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1234
					fun interpret_equal (tr1,tr2) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1235
						(case tr1 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1236
						  Leaf x =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1237
							(case tr2 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1238
							  Leaf y => PropLogic.dot_product (x,y)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1239
							| _      => raise REFUTE ("HOLogic_interpreter", "\"==\": second tree is higher"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1240
						| Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1241
							(case tr2 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1242
							  Leaf _  => raise REFUTE ("HOLogic_interpreter", "\"==\": first tree is higher")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1243
							(* extensionality: two functions are equal iff they are equal for every element *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1244
							| Node ys => PropLogic.all (map interpret_equal (xs ~~ ys))))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1245
					(* interpretation * interpretation -> prop_formula *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1246
					fun interpret_unequal (tr1,tr2) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1247
						(case tr1 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1248
						  Leaf x =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1249
							(case tr2 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1250
							  Leaf y =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1251
								let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1252
									fun unequal [] ([],_) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1253
										False
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1254
									  | unequal (x::xs) (y::ys1, ys2) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1255
										SOr (SAnd (x, PropLogic.exists (ys1@ys2)), unequal xs (ys1, y::ys2))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1256
									  | unequal _ _ =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1257
										raise REFUTE ("HOLogic_interpreter", "\"==\": leafs have different width")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1258
								in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1259
									unequal x (y,[])
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1260
								end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1261
							| _      => raise REFUTE ("HOLogic_interpreter", "\"==\": second tree is higher"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1262
						| Node xs =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1263
							(case tr2 of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1264
							  Leaf _  => raise REFUTE ("HOLogic_interpreter", "\"==\": first tree is higher")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1265
							(* extensionality: two functions are unequal iff there exist unequal elements *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1266
							| Node ys => PropLogic.exists (map interpret_unequal (xs ~~ ys))))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1267
					val fmTrue  = interpret_equal (i1,i2)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1268
					val fmFalse = interpret_unequal (i1,i2)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1269
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1270
					Some (Leaf [fmTrue, fmFalse], m2, a2)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1271
				end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1272
		| Const ("op =", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1273
			Some (interpret thy model args (eta_expand t 1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1274
		| Const ("op =", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1275
			Some (interpret thy model args (eta_expand t 2))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1276
		| Const ("op &", _) $ t1 $ t2 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1277
			apply_interpreter thy model args t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1278
		| Const ("op &", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1279
			apply_interpreter thy model args t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1280
		| Const ("op &", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1281
			Some (Node [Node [TT, FF], Node [FF, FF]], model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1282
		| Const ("op |", _) $ t1 $ t2 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1283
			apply_interpreter thy model args t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1284
		| Const ("op |", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1285
			apply_interpreter thy model args t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1286
		| Const ("op |", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1287
			Some (Node [Node [TT, TT], Node [TT, FF]], model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1288
		| Const ("op -->", _) $ t1 $ t2 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1289
			apply_interpreter thy model args t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1290
		| Const ("op -->", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1291
			apply_interpreter thy model args t
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1292
		| Const ("op -->", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1293
			Some (Node [Node [TT, FF], Node [TT, TT]], model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1294
		| Const ("Collect", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1295
			(* Collect == identity *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1296
			Some (interpret thy model args t1)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1297
		| Const ("Collect", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1298
			Some (interpret thy model args (eta_expand t 1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1299
		| Const ("op :", _) $ t1 $ t2 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1300
			(* op: == application *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1301
			Some (interpret thy model args (t2 $ t1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1302
		| Const ("op :", _) $ t1 =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1303
			Some (interpret thy model args (eta_expand t 1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1304
		| Const ("op :", _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1305
			Some (interpret thy model args (eta_expand t 2))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1306
		| _ => None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1307
	end;
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1308
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1309
	(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1310
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1311
	fun IDT_interpreter thy model args t =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1312
	let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1313
		fun is_var (Free _) = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1314
		  | is_var (Var _)  = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1315
		  | is_var _        = false
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1316
		fun typeof (Free (_,T)) = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1317
		  | typeof (Var (_,T))  = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1318
		  | typeof _            = raise REFUTE ("var_IDT_interpreter", "term is not a variable")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1319
		val (sizes, intrs) = model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1320
		(* power(a,b) computes a^b, for a>=0, b>=0 *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1321
		(* int * int -> int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1322
		fun power (a,0) = 1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1323
		  | power (a,1) = a
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1324
		  | power (a,b) = let val ab = power(a,b div 2) in ab * ab * power(a,b mod 2) end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1325
		(* interpretation -> int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1326
		fun size_of_interpretation (Leaf xs) = length xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1327
		  | size_of_interpretation (Node xs) = power (size_of_interpretation (hd xs), length xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1328
		(* Term.typ -> int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1329
		fun size_of_type T =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1330
			let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1331
				val (i,_,_) = interpret thy model args (Free ("<IDT>", T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1332
			in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1333
				size_of_interpretation i
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1334
			end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1335
		(* int list -> int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1336
		fun sum xs = foldl op+ (0, xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1337
		(* int list -> int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1338
		fun product xs = foldl op* (1, xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1339
		(* DatatypeAux.dtyp * Term.typ -> DatatypeAux.dtyp -> Term.typ *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1340
		fun typ_of_dtyp typ_assoc (DatatypeAux.DtTFree a) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1341
			the (assoc (typ_assoc, DatatypeAux.DtTFree a))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1342
		  | typ_of_dtyp typ_assoc (DatatypeAux.DtRec i) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1343
			raise REFUTE ("var_IDT_interpreter", "recursive datatype")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1344
		  | typ_of_dtyp typ_assoc (DatatypeAux.DtType (s, ds)) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1345
			Type (s, map (typ_of_dtyp typ_assoc) ds)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1346
	in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1347
		if is_var t then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1348
			(case typeof t of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1349
			  Type (s, Ts) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1350
				(case DatatypePackage.datatype_info thy s of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1351
				  Some info =>  (* inductive datatype *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1352
					let
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1353
						val index               = #index info
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1354
						val descr               = #descr info
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1355
						val (_, dtyps, constrs) = the (assoc (descr, index))
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1356
					in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1357
						if Library.exists (fn d => case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1358
							raise REFUTE ("var_IDT_interpreter", "recursive datatype argument")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1359
						else if Library.exists (fn (_,ds) => Library.exists DatatypeAux.is_rec_type ds) constrs then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1360
							None  (* recursive datatype (requires an infinite model) *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1361
						else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1362
							case assoc (intrs, t) of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1363
							  Some intr => Some (intr, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1364
							| None      =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1365
								let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1366
									val typ_assoc = dtyps ~~ Ts
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1367
									val size = sum (map (fn (_,ds) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1368
										product (map (fn d => size_of_type (typ_of_dtyp typ_assoc d)) ds)) constrs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1369
									val idx  = #next_idx args
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1370
									(* for us, an IDT has no "internal structure" -- its interpretation is just a leaf *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1371
									val intr = (if size=2 then Leaf [BoolVar idx, Not (BoolVar idx)] else Leaf (map BoolVar (idx upto (idx+size-1))))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1372
									val next = (if size=2 then idx+1 else idx+size)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1373
								in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1374
									(* extend the model, increase 'next_idx' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1375
									Some (intr, (sizes, (t, intr)::intrs), {next_idx = next, bounds = #bounds args, wellformed = #wellformed args})
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1376
								end
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1377
					end
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1378
				| None => None)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1379
			| _ => None)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1380
		else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1381
		let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1382
			val (head, params) = strip_comb t  (* look into applications to avoid unfolding *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1383
		in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1384
			(case head of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1385
			  Const (c, T) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1386
				(case body_type T of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1387
				  Type (s, Ts) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1388
					(case DatatypePackage.datatype_info thy s of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1389
					  Some info =>  (* inductive datatype *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1390
						let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1391
							val index               = #index info
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1392
							val descr               = #descr info
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1393
							val (_, dtyps, constrs) = the (assoc (descr, index))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1394
						in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1395
							if Library.exists (fn d => case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1396
								raise REFUTE ("var_IDT_interpreter", "recursive datatype argument")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1397
							else if Library.exists (fn (_,ds) => Library.exists DatatypeAux.is_rec_type ds) constrs then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1398
								None  (* recursive datatype (requires an infinite model) *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1399
							else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1400
								(case take_prefix (fn (c',_) => c' <> c) constrs of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1401
								  (_, []) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1402
									None (* unknown constructor *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1403
								| (prevs, _) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1404
									if null params then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1405
									let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1406
										val typ_assoc = dtyps ~~ Ts
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1407
										val offset = sum (map (fn (_,ds) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1408
											product (map (fn d => size_of_type (typ_of_dtyp typ_assoc d)) ds)) prevs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1409
										val (i,_,_) = interpret thy model args (Free ("<IDT>", T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1410
										(* int * interpretation  -> int * interpretation *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1411
										fun replace (offset, Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1412
											(* replace the offset-th element by True, all others by False, inc. offset by 1 *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1413
											(offset+1, Leaf (snd (foldl_map (fn (n,_) => (n+1, if n=offset then True else False)) (0, xs))))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1414
										  | replace (offset, Node xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1415
											let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1416
												val (offset', xs') = foldl_map replace (offset, xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1417
											in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1418
												(offset', Node xs')
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1419
											end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1420
										val (_,intr) = replace (offset, i)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1421
									in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1422
										Some (intr, model, args)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1423
									end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1424
									else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1425
										apply_interpreter thy model args t)  (* avoid unfolding by calling 'apply_interpreter' directly *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1426
						end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1427
					| None => None)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1428
				| _ => None)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1429
			| _ => None)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1430
		end
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1431
	end;
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1432
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1433
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1434
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1435
(* PRINTERS                                                                  *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1436
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1437
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1438
	(* theory -> model -> Term.term -> interpretation -> (int -> bool) -> string option *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1439
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1440
	fun var_typevar_printer thy model t intr assignment =
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1441
	let
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1442
		fun is_var (Free _) = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1443
		  | is_var (Var _)  = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1444
		  | is_var _        = false
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1445
		fun typeof (Free (_,T)) = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1446
		  | typeof (Var (_,T))  = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1447
		  | typeof _            = raise REFUTE ("var_typevar_printer", "term is not a variable")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1448
		fun is_typevar (TFree _) = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1449
		  | is_typevar (TVar _)  = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1450
		  | is_typevar _         = false
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1451
	in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1452
		if is_var t andalso is_typevar (typeof t) then
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1453
			let
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1454
				(* interpretation -> int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1455
				fun index_from_interpretation (Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1456
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1457
						val idx = find_index (PropLogic.eval assignment) xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1458
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1459
						if idx<0 then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1460
							raise REFUTE ("var_typevar_printer", "illegal interpretation: no value assigned")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1461
						else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1462
							idx
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1463
					end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1464
				  | index_from_interpretation _ =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1465
					raise REFUTE ("var_typevar_printer", "interpretation is not a leaf")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1466
				(* string -> string *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1467
				fun strip_leading_quote s =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1468
					(implode o (fn (x::xs) => if x="'" then xs else (x::xs)) o explode) s
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1469
				(* Term.typ -> string *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1470
				fun string_of_typ (TFree (x,_))    = strip_leading_quote x
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1471
				  | string_of_typ (TVar ((x,i),_)) = strip_leading_quote x ^ string_of_int i
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1472
				  | string_of_typ _                = raise REFUTE ("var_typevar_printer", "type is not a type variable")
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1473
			in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1474
				Some (string_of_typ (typeof t) ^ string_of_int (index_from_interpretation intr))
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1475
			end
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1476
		else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1477
			None
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1478
	end;
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1479
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1480
	(* theory -> model -> Term.term -> interpretation -> (int -> bool) -> string option *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1481
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1482
	fun var_funtype_printer thy model t intr assignment =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1483
	let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1484
		fun is_var (Free _) = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1485
		  | is_var (Var _)  = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1486
		  | is_var _        = false
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1487
		fun typeof (Free (_,T)) = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1488
		  | typeof (Var (_,T))  = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1489
		  | typeof _            = raise REFUTE ("var_funtype_printer", "term is not a variable")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1490
		fun is_funtype (Type ("fun", [_,_])) = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1491
		  | is_funtype _                     = false
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1492
	in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1493
		if is_var t andalso is_funtype (typeof t) then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1494
			let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1495
				val T1         = domain_type (typeof t)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1496
				val T2         = range_type (typeof t)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1497
				val (sizes, _) = model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1498
				(* create all constants of type T1 *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1499
				val (i, _, _) = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free ("var_funtype_printer", T1))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1500
				(* returns a list with all unit vectors of length n *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1501
				(* int -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1502
				fun unit_vectors n =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1503
				let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1504
					(* returns the k-th unit vector of length n *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1505
					(* int * int -> interpretation *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1506
					fun unit_vector (k,n) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1507
						Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1508
					(* int -> interpretation list -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1509
					fun unit_vectors_acc k vs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1510
						if k>n then [] else (unit_vector (k,n))::(unit_vectors_acc (k+1) vs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1511
				in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1512
					unit_vectors_acc 1 []
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1513
				end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1514
				(* concatenates 'x' with every list in 'xss', returning a new list of lists *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1515
				(* 'a -> 'a list list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1516
				fun cons_list x xss =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1517
					map (fn xs => x::xs) xss
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1518
				(* returns a list of lists, each one consisting of n (possibly identical) elements from 'xs' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1519
				(* int -> 'a list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1520
				fun pick_all 1 xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1521
					map (fn x => [x]) xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1522
				  | pick_all n xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1523
					let val rec_pick = pick_all (n-1) xs in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1524
						foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1525
					end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1526
				(* interpretation -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1527
				fun make_constants (Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1528
					unit_vectors (length xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1529
				  | make_constants (Node xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1530
					map (fn xs' => Node xs') (pick_all (length xs) (make_constants (hd xs)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1531
				(* interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1532
				val results = (case intr of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1533
					  Node xs => xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1534
					| _       => raise REFUTE ("var_funtype_printer", "interpretation is a leaf"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1535
				(* string list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1536
				val strings = map
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1537
					(fn (argi,resulti) => print thy model (Free ("var_funtype_printer", T1)) argi assignment
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1538
						^ "\\<mapsto>"
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1539
						^ print thy model (Free ("var_funtype_printer", T2)) resulti assignment)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1540
					((make_constants i) ~~ results)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1541
			in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1542
				Some (enclose "(" ")" (commas strings))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1543
			end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1544
		else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1545
			None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1546
	end;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1547
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1548
	(* theory -> model -> Term.term -> interpretation -> (int -> bool) -> string option *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1549
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1550
	fun var_settype_printer thy model t intr assignment =
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1551
	let
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1552
		val (sizes, _) = model
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1553
		(* returns a list with all unit vectors of length n *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1554
		(* int -> interpretation list *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1555
		fun unit_vectors n =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1556
		let
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1557
			(* returns the k-th unit vector of length n *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1558
			(* int * int -> interpretation *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1559
			fun unit_vector (k,n) =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1560
				Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1561
			(* int -> interpretation list -> interpretation list *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1562
			fun unit_vectors_acc k vs =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1563
				if k>n then [] else (unit_vector (k,n))::(unit_vectors_acc (k+1) vs)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1564
		in
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1565
			unit_vectors_acc 1 []
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1566
		end
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1567
		(* concatenates 'x' with every list in 'xss', returning a new list of lists *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1568
		(* 'a -> 'a list list -> 'a list list *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1569
		fun cons_list x xss =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1570
			map (fn xs => x::xs) xss
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1571
		(* returns a list of lists, each one consisting of n (possibly identical) elements from 'xs' *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1572
		(* int -> 'a list -> 'a list list *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1573
		fun pick_all 1 xs =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1574
			map (fn x => [x]) xs
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1575
		  | pick_all n xs =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1576
			let val rec_pick = pick_all (n-1) xs in
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1577
				foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1578
			end
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1579
		(* interpretation -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1580
		fun make_constants (Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1581
			unit_vectors (length xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1582
		  | make_constants (Node xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1583
			map (fn xs' => Node xs') (pick_all (length xs) (make_constants (hd xs)))
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1584
	in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1585
		case t of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1586
		  Free (x, Type ("set", [T])) =>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1587
			let
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1588
				(* interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1589
				val results = (case intr of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1590
				  Node xs => xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1591
					| _       => raise REFUTE ("var_settype_printer", "interpretation is a leaf"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1592
				(* create all constants of type T *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1593
				val (i, _, _) = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free ("var_settype_printer", T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1594
				(* string list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1595
				val strings = mapfilter
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1596
					(fn (argi,Leaf [fmTrue,fmFalse]) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1597
						if PropLogic.eval assignment fmTrue then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1598
							Some (print thy model (Free ("x", T)) argi assignment)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1599
						else if PropLogic.eval assignment fmFalse then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1600
							None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1601
						else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1602
							raise REFUTE ("var_settype_printer", "illegal interpretation: no value assigned"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1603
					((make_constants i) ~~ results)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1604
			in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1605
				Some (enclose "{" "}" (commas strings))
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1606
			end
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1607
		| Var ((x,i), Type ("set", [T])) =>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1608
			let
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1609
				(* interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1610
				val results = (case intr of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1611
				  Node xs => xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1612
					| _       => raise REFUTE ("var_settype_printer", "interpretation is a leaf"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1613
				(* create all constants of type T *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1614
				val (i, _, _) = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free ("var_settype_printer", T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1615
				(* string list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1616
				val strings = mapfilter
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1617
					(fn (argi,Leaf [fmTrue,fmFalse]) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1618
						if PropLogic.eval assignment fmTrue then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1619
							Some (print thy model (Free ("var_settype_printer", T)) argi assignment)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1620
						else if PropLogic.eval assignment fmTrue then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1621
							None
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1622
						else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1623
							raise REFUTE ("var_settype_printer", "illegal interpretation: no value assigned"))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1624
					((make_constants i) ~~ results)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1625
			in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1626
				Some (enclose "{" "}" (commas strings))
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1627
			end
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1628
		| _ => None
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1629
	end;
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1630
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1631
	(* theory -> model -> Term.term -> interpretation -> (int -> bool) -> string option *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1632
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1633
	fun HOLogic_printer thy model t intr assignment =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1634
		case t of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1635
		(* 'arbitrary', 'The', 'Hilbert_Choice.Eps' are printed like free variables of the same type *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1636
		  Const ("arbitrary", T) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1637
			Some (print thy model (Free ("<arbitrary>", T)) intr assignment)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1638
		| Const ("The", T) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1639
			Some (print thy model (Free ("<The>", T)) intr assignment)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1640
		| Const ("Hilbert_Choice.Eps", T) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1641
			Some (print thy model (Free ("<Eps>", T)) intr assignment)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1642
		| _ =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1643
			None;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1644
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1645
	(* theory -> model -> Term.term -> interpretation -> (int -> bool) -> string option *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1646
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1647
	fun var_IDT_printer thy model t intr assignment =
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1648
	let
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1649
		fun is_var (Free _) = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1650
		  | is_var (Var _)  = true
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1651
		  | is_var _        = false
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1652
		fun typeof (Free (_,T)) = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1653
		  | typeof (Var (_,T))  = T
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1654
		  | typeof _            = raise REFUTE ("var_IDT_printer", "term is not a variable")
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1655
	in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1656
		if is_var t then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1657
			(case typeof t of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1658
			  Type (s, Ts) =>
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1659
				(case DatatypePackage.datatype_info thy s of
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1660
				  Some info =>  (* inductive datatype *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1661
					let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1662
						val index               = #index info
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1663
						val descr               = #descr info
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1664
						val (_, dtyps, constrs) = the (assoc (descr, index))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1665
					in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1666
						if Library.exists (fn d => case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1667
							raise REFUTE ("var_IDT_printer", "recursive datatype argument")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1668
						else if Library.exists (fn (_,ds) => Library.exists DatatypeAux.is_rec_type ds) constrs then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1669
							None  (* recursive datatype (requires an infinite model) *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1670
						else
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1671
						let
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1672
							val (sizes, _) = model
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1673
							val typ_assoc  = dtyps ~~ Ts
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1674
							(* interpretation -> int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1675
							fun index_from_interpretation (Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1676
								let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1677
									val idx = find_index (PropLogic.eval assignment) xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1678
								in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1679
									if idx<0 then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1680
										raise REFUTE ("var_IDT_printer", "illegal interpretation: no value assigned")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1681
									else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1682
										idx
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1683
								end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1684
							  | index_from_interpretation _ =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1685
								raise REFUTE ("var_IDT_printer", "interpretation is not a leaf")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1686
							(* string -> string *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1687
							fun unqualify s =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1688
								implode (snd (take_suffix (fn c => c <> ".") (explode s)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1689
							(* DatatypeAux.dtyp -> Term.typ *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1690
							fun typ_of_dtyp (DatatypeAux.DtTFree a) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1691
								the (assoc (typ_assoc, DatatypeAux.DtTFree a))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1692
							  | typ_of_dtyp (DatatypeAux.DtRec i) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1693
								raise REFUTE ("var_IDT_printer", "recursive datatype")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1694
							  | typ_of_dtyp (DatatypeAux.DtType (s, ds)) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1695
								Type (s, map typ_of_dtyp ds)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1696
							fun sum xs     = foldl op+ (0, xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1697
							fun product xs = foldl op* (1, xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1698
							(* power(a,b) computes a^b, for a>=0, b>=0 *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1699
							(* int * int -> int *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1700
							fun power (a,0) = 1
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1701
							  | power (a,1) = a
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1702
							  | power (a,b) = let val ab = power(a,b div 2) in ab * ab * power(a,b mod 2) end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1703
							fun size_of_interpretation (Leaf xs) = length xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1704
							  | size_of_interpretation (Node xs) = power (size_of_interpretation (hd xs), length xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1705
							fun size_of_type T =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1706
								let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1707
									val (i,_,_) = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free ("<IDT>", T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1708
								in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1709
									size_of_interpretation i
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1710
								end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1711
							(* returns a list with all unit vectors of length n *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1712
							(* int -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1713
							fun unit_vectors n =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1714
							let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1715
								(* returns the k-th unit vector of length n *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1716
								(* int * int -> interpretation *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1717
								fun unit_vector (k,n) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1718
									Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1719
								(* int -> interpretation list -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1720
								fun unit_vectors_acc k vs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1721
									if k>n then [] else (unit_vector (k,n))::(unit_vectors_acc (k+1) vs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1722
							in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1723
								unit_vectors_acc 1 []
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1724
							end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1725
							(* concatenates 'x' with every list in 'xss', returning a new list of lists *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1726
							(* 'a -> 'a list list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1727
							fun cons_list x xss =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1728
								map (fn xs => x::xs) xss
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1729
							(* returns a list of lists, each one consisting of n (possibly identical) elements from 'xs' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1730
							(* int -> 'a list -> 'a list list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1731
							fun pick_all 1 xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1732
								map (fn x => [x]) xs
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1733
							  | pick_all n xs =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1734
								let val rec_pick = pick_all (n-1) xs in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1735
									foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1736
								end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1737
							(* interpretation -> interpretation list *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1738
							fun make_constants (Leaf xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1739
								unit_vectors (length xs)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1740
							  | make_constants (Node xs) =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1741
								map (fn xs' => Node xs') (pick_all (length xs) (make_constants (hd xs)))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1742
							(* DatatypeAux.dtyp list -> int -> string *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1743
							fun string_of_inductive_type_cargs [] n =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1744
								if n<>0 then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1745
									raise REFUTE ("var_IDT_printer", "internal error computing the element index for an inductive type")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1746
								else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1747
									""
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1748
							  | string_of_inductive_type_cargs (d::ds) n =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1749
								let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1750
									val size_ds   = product (map (fn d => size_of_type (typ_of_dtyp d)) ds)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1751
									val T         = typ_of_dtyp d
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1752
									val (i,_,_)   = interpret thy (sizes, []) {next_idx=1, bounds=[], wellformed=True} (Free ("<IDT>", T))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1753
									val constants = make_constants i
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1754
								in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1755
									" "
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1756
									^ (print thy model (Free ("<IDT>", T)) (nth_elem (n div size_ds, constants)) assignment)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1757
									^ (string_of_inductive_type_cargs ds (n mod size_ds))
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1758
								end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1759
							(* (string * DatatypeAux.dtyp list) list -> int -> string *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1760
							fun string_of_inductive_type_constrs [] n =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1761
								raise REFUTE ("var_IDT_printer", "inductive type has fewer elements than needed")
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1762
							  | string_of_inductive_type_constrs ((c,ds)::cs) n =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1763
								let
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1764
									val size = product (map (fn d => size_of_type (typ_of_dtyp d)) ds)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1765
								in
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1766
									if n < size then
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1767
										(unqualify c) ^ (string_of_inductive_type_cargs ds n)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1768
									else
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1769
										string_of_inductive_type_constrs cs (n - size)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1770
								end
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1771
						in
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1772
							Some (string_of_inductive_type_constrs constrs (index_from_interpretation intr))
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1773
						end
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1774
					end
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1775
				| None => None)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1776
			| _ => None)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1777
		else
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1778
			None
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1779
	end;
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1780
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1781
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1782
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1783
(* use 'setup Refute.setup' in an Isabelle theory to initialize the 'Refute' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1784
(* structure                                                                 *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1785
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1786
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1787
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1788
(* Note: the interpreters and printers are used in reverse order; however,   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1789
(*       an interpreter that can handle non-atomic terms ends up being       *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1790
(*       applied before other interpreters are applied to subterms!          *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1791
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1792
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1793
	(* (theory -> theory) list *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1794
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1795
	val setup =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1796
		[RefuteData.init,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1797
		 add_interpreter "var_typevar"   var_typevar_interpreter,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1798
		 add_interpreter "var_funtype"   var_funtype_interpreter,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1799
		 add_interpreter "var_settype"   var_settype_interpreter,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1800
		 add_interpreter "boundvar"      boundvar_interpreter,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1801
		 add_interpreter "abstraction"   abstraction_interpreter,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1802
		 add_interpreter "apply"         apply_interpreter,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1803
		 add_interpreter "const_unfold"  const_unfold_interpreter,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1804
		 add_interpreter "Pure"          Pure_interpreter,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1805
		 add_interpreter "HOLogic"       HOLogic_interpreter,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1806
		 add_interpreter "IDT"           IDT_interpreter,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1807
		 add_printer "var_typevar"   var_typevar_printer,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1808
		 add_printer "var_funtype"   var_funtype_printer,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1809
		 add_printer "var_settype"   var_settype_printer,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1810
		 add_printer "HOLogic"       HOLogic_printer,
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1811
		 add_printer "var_IDT"       var_IDT_printer];
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1812
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1813
end