Theory Simple_Network_Language_Certificate_Code

section ‹Assembling the Checker and Generating Code›

text ‹\label{export-muntac}
We provide some optimized code equations.

Then we assemble the certificate checker:

 A parser is added to parse JSON files
 The resulting JSON data is validated and converted to the simple networks language
 The (binary) certificate is read
 There is an additional ``renaming'' that maps identifiers between the input and the certificate
 The certficiate, the renaming, and the model are passed to the checker
 and the output is printed

Then the code is exported:

 We add some logging utilities (via code printings)
 Code generation is setup
 We export this to SML via reflection
 And test it on a few small examples
›

theory Simple_Network_Language_Certificate_Code
  imports
    Simple_Network_Language_Certificate
    Simple_Network_Language_Certificate_Checking
begin

paragraph ‹Optimized code equations›

lemmas [code_unfold] = imp_for_imp_for'

definition dbm_subset'_impl'_int
  :: "nat  int DBMEntry Heap.array  int DBMEntry Heap.array  bool Heap"
  where [symmetric, int_folds]:
    "dbm_subset'_impl'_int = dbm_subset'_impl'"

lemma less_eq_dbm_le_int[int_folds]:
  "(≤) = dbm_le_int"
  unfolding dbm_le_dbm_le_int less_eq ..

schematic_goal dbm_subset'_impl'_int_code[code]:
  "dbm_subset'_impl'_int  λm a b.
    do {
    imp_for 0 ((m + 1) * (m + 1)) Heap_Monad.return
      (λi _. do {
        x  Array.nth a i; y  Array.nth b i; Heap_Monad.return (dbm_le_int x y)
      })
      True
    }
"
  unfolding dbm_subset'_impl'_int_def[symmetric] dbm_subset'_impl'_def int_folds .


definition dbm_subset_impl_int
  :: "nat  int DBMEntry Heap.array  int DBMEntry Heap.array  bool Heap"
  where [symmetric, int_folds]:
    "dbm_subset_impl_int = dbm_subset_impl"

schematic_goal dbm_subset_impl_int_code[code]:
  "dbm_subset_impl_int  ?i"
  unfolding dbm_subset_impl_int_def[symmetric] dbm_subset_impl_def
  unfolding int_folds
  .

lemmas [code_unfold] = int_folds

export_code state_space in SML module_name Test

hide_const Parser_Combinator.return
hide_const Error_Monad.return

definition
  "show_lit = String.implode o show"

definition "rename_state_space  λdc ids_to_names (broadcast, automata, bounds) L0 s0 formula.
  let _ = println (STR ''Make renaming'') in
  do {
    (m, num_states, num_actions, renum_acts, renum_vars, renum_clocks, renum_states,
      inv_renum_states, inv_renum_vars, inv_renum_clocks)
       make_renaming broadcast automata bounds;
    let _ = println (STR ''Renaming'');
    let (broadcast', automata', bounds') = rename_network
      broadcast bounds automata renum_acts renum_vars renum_clocks renum_states;
    let _ = println (STR ''Calculating ceiling'');
    let k = Simple_Network_Impl_nat_defs.local_ceiling broadcast' bounds' automata' m num_states;
    let _ = println (STR ''Running model checker'');
    let inv_renum_states' = (λi. ids_to_names i o inv_renum_states i);
    let f = (λshow_clock
      show_state broadcast bounds' automata m num_states num_actions k L0 s0 formula.
      state_space broadcast bounds' automata m num_states num_actions k L0 s0 formula
        show_clock show_state ()
    );
    let r = do_rename_mc f dc broadcast bounds automata k STR ''_urge'' L0 s0 formula
      m num_states num_actions renum_acts renum_vars renum_clocks renum_states
      inv_renum_states' inv_renum_vars inv_renum_clocks;
    let show_clock = show o inv_renum_clocks;
    let show_state = (show_state :: _  _  _ × int list  _) inv_renum_states inv_renum_vars;
    let renamings =
      (m, num_states, num_actions, renum_acts, renum_vars, renum_clocks, renum_states,
       inv_renum_states, inv_renum_vars, inv_renum_clocks
      );
    Result (r, show_clock, show_state, renamings, k)
  }"

definition
  "check_subsumed n xs (i :: int) M 
  do {
    (_, b)  imp_nfoldli xs (λ(_, b). return (¬ b)) (λM' (j, b).
      if i = j then return (j + 1, b) else do {
        b  dbm_subset'_impl n M M';
        if b then return (j, True) else return (j + 1, False)
      }
    ) (0, False);
    return b
  }
"

definition
  "imp_filter_index P xs = do {
  (_, xs)  imp_nfoldli xs (λ_. return True) (λx (i :: nat, xs).
    do {
      b  P i x;
      return (i + 1, if b then (x # xs) else xs)
    }
  ) (0, []);
  return (rev xs)
  }"

definition
  "filter_dbm_list n xs =
    imp_filter_index (λi M. do {b  check_subsumed n xs i M; return (¬ b)}) xs"

partial_function (heap) imp_map :: "('a  'b Heap)  'a list  'b list Heap" where
  "imp_map f xs =
  (if xs = [] then return [] else do {y  f (hd xs); ys  imp_map f (tl xs); return (y # ys)})"

lemma imp_map_simps[code, simp]:
  "imp_map f [] = return []"
  "imp_map f (x # xs) = do {y  f x; ys  imp_map f xs; return (y # ys)}"
  by (simp add: imp_map.simps)+

definition trace_state where
  "trace_state n show_clock show_state 
  λ (l, M). do {
      let st = show_state l;
      m  show_dbm_impl n show_clock show M;
      let s = ''(''  @ st @ '', ['' @ m @ ''])''; 
      let s = String.implode s;
      let _ = println s;
      return ()
  }
" for show_clock show_state

definition
  "show_str = String.implode o show"

definition parse_convert_run_print where
  "parse_convert_run_print dc s 
   case parse json s  convert of
     Error es  do {let _ = map println es; return ()}
   | Result (ids_to_names, _, broadcast, automata, bounds, formula, L0, s0)  do {
      let r = rename_state_space dc ids_to_names (broadcast, automata, bounds) L0 s0 formula;
      case r of
        Error es  do {let _ = map println es; return ()}
      | Result (r, show_clk, show_st, renamings, k) 
        case r of None  return () | Some r 
        do {
          r  r;
          let _ = STR ''Number of discrete states: '' + (length r |> show_str) |> println;
          let _ =
            STR ''Size of passed list: '' + show_str (sum_list (map (length o snd) r)) |> println;
          let n = Simple_Network_Impl.clk_set' automata |> list_of_set |> length;
          r  imp_map (λ (a, b). do {
              b  imp_map (return o snd) b; b  filter_dbm_list n b; return (a, b)
            }) r;
          let _ = STR ''Number of discrete states: '' + show_str (length r) |> println;
          let _ = STR ''Size of passed list after removing subsumed states: ''
            + show_str (sum_list (map (length o snd) r)) |> println;
          let show_dbm = (λM. do {
            s  show_dbm_impl_all n show_clk show M;
            return (''<'' @ s @ ''>'')
          });
          _  imp_map (λ (s, xs).
          do {
            let s = show_st s;
            xs  imp_map show_dbm xs;
            let _ = s @ '': '' @ show xs |> String.implode |> println;
            return ()
          }
          ) r;
          return ()
        }
  }"

ML structure Timing : sig
  val start_timer: unit -> unit
  val save_time: string -> unit
  val get_timings: unit -> (string * Time.time) list
end = struct
  val t = Unsynchronized.ref Time.zeroTime;
  val timings = Unsynchronized.ref [];
  fun start_timer () = (t := Time.now ());
  fun get_elapsed () = Time.- (Time.now (), !t);
  fun save_time s = (timings := ((s, get_elapsed ()) :: !timings));
  fun get_timings () = !timings;
end

ML fun print_timings () =
  let
    val tab = Timing.get_timings ();
    fun print_time (s, t) = writeln(s ^ ": " ^ Time.toString t);
  in map print_time tab end;

code_printing
  constant "Show_State_Defs.tracei" 
      (SML)   "(fn n => fn show_state => fn show_clock => fn typ => fn x => ()) _ _ _"
  and (OCaml) "(fun n show_state show_clock ty x -> ()) _ _ _"

datatype mode = Impl1 | Impl2 | Impl3 | Buechi | Debug

definition
  "distr xs 
  let (m, d) =
  fold
    (λx (m, d). case m x of None  (m(x  1:: nat), x # d) | Some y  (m(x  (y + 1)), d))
    xs (Map.empty, [])
  in map (λx. (x, the (m x))) (sort d)"

definition split_k'' :: "nat  ('a × 'b list) list  ('a × 'b list) list list" where
  "split_k'' k xs  let
    width = sum_list (map (length o snd) xs) div k;
    width = (if length xs mod k = 0 then width else width + 1)
  in split_size (length o snd) width 0 [] xs"

definition
  "print_errors es = do {Heap_Monad.fold_map print_line_impl es; return ()}"

text ‹\label{self-check} We can actually let \munta\ produce reachability certificates
for \muntac. This is what the following code does.
A few test cases are run below right in the Isabelle/ML environment.›

definition parse_convert_run_check where
  "parse_convert_run_check mode num_split dc s 
   case parse json s  convert of
     Error es  print_errors es
   | Result (ids_to_names, _, broadcast, automata, bounds, formula, L0, s0)  do {
      let r = rename_state_space dc ids_to_names (broadcast, automata, bounds) L0 s0 formula;
      case r of
        Error es  print_errors es
      | Result (r, show_clk, show_st, renamings, k) 
        case r of None  return () | Some r  do {
        let t = now ();
        r  r;
        let t = now () - t;
        print_line_impl
          (STR ''Time for model checking + certificate extraction: '' + time_to_string t);
        let (m,num_states,num_actions,renum_acts,renum_vars,renum_clocks,renum_states,
          inv_renum_states, inv_renum_vars, inv_renum_clocks
        ) = renamings;
        let _ = start_timer ();
        state_space  Heap_Monad.fold_map (λ(s, xs).
          do {
            let xs = map snd xs;
            xs  Heap_Monad.fold_map (dbm_to_list_impl m) xs;
            return (s, xs)
          }
        ) r;
        let _ = save_time STR ''Time for converting DBMs in certificate'';
        print_line_impl
          (STR ''Number of discrete states of state space: '' + show_lit (length state_space));
        let _ = STR ''Size of passed list: ''  + show_str (sum_list (map (length o snd) r))
          |> println;
        STR ''DBM list length distribution: '' + show_str (distr (map (length o snd) state_space))
          |> print_line_impl;
        let split =
          (if mode = Impl3 then split_k'' num_split state_space else split_k num_split state_space);
        let split_distr = map (sum_list o map (length o snd)) split;
        STR ''Size of passed list distribution after split: '' + show_str split_distr
          |> print_line_impl;
        let t = now ();
        check  case mode of
          Debug  rename_check_dbg num_split dc broadcast bounds automata k L0 s0 formula
            m num_states num_actions renum_acts renum_vars renum_clocks renum_states
            inv_renum_states inv_renum_vars inv_renum_clocks
            state_space
        | Impl1  rename_check num_split dc broadcast bounds automata k L0 s0 formula
            m num_states num_actions renum_acts renum_vars renum_clocks renum_states
            state_space
        | Impl2  rename_check2 num_split dc broadcast bounds automata k L0 s0 formula
            m num_states num_actions renum_acts renum_vars renum_clocks renum_states
            state_space |> return
        | Impl3  rename_check3 num_split dc broadcast bounds automata k L0 s0 formula
            m num_states num_actions renum_acts renum_vars renum_clocks renum_states
            state_space |> return;
        let t = now () - t;
        print_line_impl (STR ''Time for certificate checking: '' + time_to_string t);
        case check of
          Renaming_Failed  print_line_impl (STR ''Renaming failed'')
        | Preconds_Unsat  print_line_impl (STR ''Preconditions were not met'')
        | Sat  print_line_impl (STR ''Certificate was accepted'')
        | Unsat  print_line_impl (STR ''Certificate was rejected'')
        }
    }"

ML fun do_test dc file =
  let
    val dir = @{master_dir} + @{path "benchmarks/"}
    val file = dir + Path.explode file
    val s = File.read file
  in
    @{code parse_convert_run_print} dc s end

ML_val do_test true "HDDI_02.muntax" ()

ML_val do_test true "simple.muntax" ()

ML_val do_test true "light_switch.muntax" ()

ML_val do_test true "PM_test.muntax" ()

ML_val do_test true "bridge.muntax" ()

code_printing
  constant "parallel_fold_map" 
      (SML)   "(fn f => fn xs => fn () => Par'_List.map (fn x => f x ()) xs) _ _"

definition
  "num_split  4 :: nat"

ML fun do_check dc file =
  let
    val dir = @{master_dir} + @{path "benchmarks/"}
    val file = dir + Path.explode file
    val s = File.read file
  in
    @{code parse_convert_run_check} @{code Impl3} @{code num_split} dc s end

ML_val do_check false "HDDI_02.muntax" ()

ML_val do_check true "HDDI_02.muntax" ()

ML_val do_check true "simple.muntax" ()

ML_val do_check true "light_switch.muntax" ()

ML_val do_check false "PM_test.muntax" ()

ML_val do_check true "PM_test.muntax" ()

ML_val do_check true "bridge.muntax" ()

ML_val do_check false "PM_all_3.muntax" ()

ML_val do_check true "PM_all_3.muntax" ()

text ‹Executing Heap_Monad.fold_map› in parallel in Isabelle/ML›

definition
  Test  Heap_Monad.fold_map (λx. do {let _ = println STR ''x''; return x}) ([1,2,3] :: nat list)

ML_val @{code Test} ()


paragraph ‹Checking external certificates›

definition
  "list_of_json_object obj 
  case obj of
    Object m  Result m
  | _  Error [STR ''Not an object'']
"

definition
  "map_of_debug prefix m 
    let m = map_of m in
    (λx.
      case m x of
        None  let _ = println (STR ''Key error('' + prefix + STR ''): '' + show_lit x) in None
      | Some v  Some v)" for prefix

definition
  "renaming_of_json' opt prefix json  do {
    vars  list_of_json_object json;
    vars  combine_map (λ (a, b). do {b  of_nat b; Result (String.implode a, b)}) vars;
    let vars = vars @ (case opt of Some x  [(x, fold max (map snd vars) 0 + 1)] | _  []);
    Result (the o map_of_debug prefix vars, the o map_of_debug prefix (map prod.swap vars))
  }
  " for json prefix

definition "renaming_of_json  renaming_of_json' None"

definition
  "nat_renaming_of_json prefix max_id json  do {
    vars  list_of_json_object json;
    vars  combine_map (λ(a, b). do {
      a  parse lx_nat (String.implode a);
      b  of_nat b;
      Result (a, b)
    }) vars;
    let ids = fst ` set vars;
    let missing = filter (λi. i  ids) [0..<max_id];
    let m = the o map_of_debug prefix vars;
    let m = extend_domain m missing (length vars);
    Result (m, the o map_of_debug prefix (map prod.swap vars))
  }
  " for json prefix

definition convert_renaming ::
  "(nat  nat  String.literal)  (String.literal  nat)  JSON  _" where
  "convert_renaming ids_to_names process_names_to_index json  do {
    json  of_object json;
    vars  get json ''vars'';
    (var_renaming, var_inv)  renaming_of_json STR ''var renaming'' vars;
    clocks  get json ''clocks'';
    (clock_renaming, clock_inv)
       renaming_of_json' (Some STR ''_urge'') STR ''clock renaming'' clocks;
    processes  get json ''processes'';
    (process_renaming, process_inv)  renaming_of_json STR ''process renaming'' processes;
    locations  get json ''locations'';
    locations  list_of_json_object locations; ―‹process name →› json›
    locations  combine_map (λ(name, renaming). do {
        let p_num = process_names_to_index (String.implode name);
        assert
          (process_renaming (String.implode name) = p_num)
          (STR ''Process renamings do not agree on '' + String.implode name);
        let max_id = 1000;
        renaming  nat_renaming_of_json (STR ''process'' + show_str p_num) max_id renaming;
          ―‹location id →› nat›
        Result (p_num, renaming)
      }
      ) locations;
    ―‹process id →› location id →› nat›
    let location_renaming = the o map_of_debug STR ''location'' (map (λ(i, f, _). (i, f)) locations);
    let location_inv = the o map_of_debug STR ''location inv''  (map (λ(i, _, f). (i, f)) locations);
    Result (var_renaming, clock_renaming, location_renaming, var_inv, clock_inv, location_inv)
  }
  "
  for json

definition
  "load_renaming dc model renaming 
  case
  do {
    model  parse json model;
    renaming  parse json renaming;
    (ids_to_names, process_names_to_index, broadcast, automata, bounds, formula, L0, s0)
       convert model;
    convert_renaming ids_to_names process_names_to_index renaming
  }
  of
    Error e  return (Error e)
  | Result r  do {
    let (var_renaming, clock_renaming, location_renaming, _, _, _) = r;
    let _ = map (λp. map (λn. location_renaming p n |> show_lit |> println) [0..<8]) [0..<6];
    return (Result ())
  }
"

definition parse_compute where
  "parse_compute model renaming 
   do {
    model  parse json model;
    (ids_to_names, process_names_to_index, broadcast, automata, bounds, formula, L0, s0)
       convert model;
    renaming  parse json renaming;
    (var_renaming, clock_renaming, location_renaming,
     inv_renum_vars, inv_renum_clocks, inv_renum_states)
       convert_renaming ids_to_names process_names_to_index renaming;
    (m, num_states, num_actions, renum_acts, renum_vars, renum_clocks, renum_states, _, _, _)
       make_renaming broadcast automata bounds;
    assert (renum_clocks STR ''_urge'' = m) STR ''Computed renaming: _urge is not last clock!'';
    let renum_vars = var_renaming;
    let renum_clocks = clock_renaming;
    let renum_states = location_renaming;
    assert (renum_clocks STR ''_urge'' = m) STR ''Given renaming: _urge is not last clock!'';
    let _ = println (STR ''Renaming'');
    let (broadcast', automata', bounds') = rename_network
      broadcast bounds automata renum_acts renum_vars renum_clocks renum_states;
    let _ = println (STR ''Calculating ceiling'');
    let k = Simple_Network_Impl_nat_defs.local_ceiling broadcast' bounds' automata' m num_states;
    let urgent_locations = map (λ(_, urgent, _, _). urgent) automata';
    Result (broadcast, bounds, automata, urgent_locations, k, L0, s0, formula,
          m, num_states, num_actions, renum_acts, renum_vars, renum_clocks, renum_states,
          inv_renum_states, inv_renum_vars, inv_renum_clocks)
   }" for num_split

datatype 'l state_space =
    Reachable_Set (reach_of: "(('l list × int list) × int DBMEntry list list) list")
  | Buechi_Set    (buechi_of: "(('l list × int list) × (int DBMEntry list × nat) list) list")

definition normalize_dbm :: "nat  int DBMEntry list  int DBMEntry list" where
  "normalize_dbm m xs = do {
    dbm  Array.of_list xs;
    dbm  fw_impl_int m dbm;
    Array.freeze dbm
  } |> run_heap
  "

definition insert_every_nth :: "nat  'a  'a list  'a list" where
  "insert_every_nth n a xs 
    fold (λx (i, xs). if i = n then (1, a # x # xs) else (i + 1, x # xs)) xs (1, []) |> snd |> rev"

definition convert_dbm where
  "convert_dbm urge m dbm =
    take m dbm @ Le 0 #
    insert_every_nth m DBMEntry.INF (drop m dbm)
    @ (if urge then Le 0 else DBMEntry.INF) # replicate (m - 1) DBMEntry.INF @ [Le 0]
  |> normalize_dbm m"

fun convert_state_space :: "_  _  int state_space  nat state_space" where
  "convert_state_space m is_urgent (Reachable_Set xs) = Reachable_Set (
    map (λ((locs, vars), dbms).
      ((map nat locs, vars), map (convert_dbm (is_urgent (locs, vars)) m) dbms))
    xs)"
| "convert_state_space m is_urgent (Buechi_Set xs) = Buechi_Set (
    map (λ((locs, vars), dbms).
      ((map nat locs, vars), map (λ(dbm, i). (convert_dbm (is_urgent (locs, vars)) m dbm, i)) dbms))
    xs)"

fun len_of_state_space where
  "len_of_state_space (Reachable_Set xs) = length xs"
| "len_of_state_space (Buechi_Set xs) = length xs"

context
  fixes num_clocks :: nat
  fixes inv_renum_states :: "nat  nat  nat"
    and inv_renum_vars   :: "nat  String.literal"
    and inv_renum_clocks :: "nat  String.literal"
begin

private definition show_st where
  "show_st = show_state inv_renum_states inv_renum_vars"

private definition show_dbm :: "int DBMEntry list  char list" where
  "show_dbm = dbm_list_to_string num_clocks (show_clock inv_renum_clocks) show"

fun show_state_space where
  "show_state_space (Reachable_Set xs) =
    map (λ(l, xs).
      map (λx. ''('' @ show_st l @ '', <'' @ show_dbm x @ ''>)'' |> String.implode |> println) xs)
    xs"
| "show_state_space (Buechi_Set xs) =
    map (λ(l, xs).
      map (λ(x, i). show i @ '': ('' @ show_st l @ '', <'' @ show_dbm x @ ''>)''
            |> String.implode |> println)
      xs)
    xs"

end

definition
  "print_sep  λ(). println (String.implode (replicate 100 CHR ''-''))"

definition parse_convert_check where
  "parse_convert_check mode num_split dc model renaming state_space show_cert 
   let
     r = parse_compute model renaming
   in case r of Error es  do {let _ = map println es; return ()}
   | Result r  do {
      let (broadcast, bounds, automata, urgent_locations, k, L0, s0, formula,
        m, num_states, num_actions, renum_acts, renum_vars, renum_clocks, renum_states,
        inv_renum_states, inv_renum_vars, inv_renum_clocks) = r;
      let is_urgent = (λ(L, _). list_ex (λ(l, urgent). l  set urgent) (zip L urgent_locations));
      let inv_renum_clocks = (λi. if i = m then STR ''_urge'' else inv_renum_clocks i);
      let t = now ();
      let state_space = convert_state_space m is_urgent state_space;
      let t = now () - t;
      let _ = println (STR ''Time for converting state space: '' + time_to_string t);
      let _ = start_timer ();
      let _ = save_time STR ''Time for converting DBMs in certificate'';
      let _ =
        println (STR ''Number of discrete states: ''+ show_lit (len_of_state_space state_space));
      let _ = do {
        if show_cert then do {
          let _ = print_sep ();
          let _ = println (STR ''Certificate'');
          let _ = print_sep ();
          let _ = show_state_space m inv_renum_states inv_renum_vars inv_renum_clocks state_space;
          let _ = print_sep ();
          return ()}
        else return ()
      };
      let t = now ();
      check  case mode of
        Debug  rename_check_dbg num_split dc broadcast bounds automata k L0 s0 formula
            m num_states num_actions renum_acts renum_vars renum_clocks renum_states
            inv_renum_states inv_renum_vars inv_renum_clocks
            (reach_of state_space)
      | Impl1  rename_check num_split dc broadcast bounds automata k L0 s0 formula
          m num_states num_actions renum_acts renum_vars renum_clocks renum_states
          (reach_of state_space)
      | Impl2  rename_check2 num_split dc broadcast bounds automata k L0 s0 formula
          m num_states num_actions renum_acts renum_vars renum_clocks renum_states
          (reach_of state_space) |> return
      | Impl3  rename_check3 num_split dc broadcast bounds automata k L0 s0 formula
          m num_states num_actions renum_acts renum_vars renum_clocks renum_states
          (reach_of state_space) |> return
      | Buechi  rename_check_buechi num_split broadcast bounds automata k L0 s0 formula
          m num_states num_actions renum_acts renum_vars renum_clocks renum_states
          (buechi_of state_space) |> return;
      let t = now () - t;
      let _ = println (STR ''Time for certificate checking: '' + time_to_string t);
      case check of
        Renaming_Failed  do {let _ = println STR ''Renaming failed''; return ()}
      | Preconds_Unsat  do {let _ = println STR ''Preconditions were not met''; return ()}
      | Sat  do {let _ = println STR ''Certificate was accepted''; return ()}
      | Unsat  do {let _ = println STR ''Certificate was rejected''; return ()}
    }
" for num_split and state_space :: "int state_space"

(* XXX This is a bug fix. Should be fixed in the Isabelle distribution. *)
code_printing
  constant IArray.length'  (SML) "(IntInf.fromInt o Vector.length)"

code_printing
  constant Parallel.map  (SML) "Par'_List.map"

lemma [code]: "run_map_heap f xs = Parallel.map (run_heap o f) xs"
  unfolding run_map_heap_def Parallel.map_def ..

code_printing code_module "Timing"  (SML)
‹
structure Timing : sig
  val start_timer: unit -> unit
  val save_time: string -> unit
  val get_timings: unit -> (string * Time.time) list
  val set_cpu: bool -> unit
end = struct

  open Timer;

  val is_cpu = Unsynchronized.ref false;
  fun set_cpu b = is_cpu := b;

  val cpu_timer: cpu_timer option Unsynchronized.ref = Unsynchronized.ref NONE;
  val real_timer: real_timer option Unsynchronized.ref = Unsynchronized.ref NONE;

  val timings = Unsynchronized.ref [];
  fun start_timer () = (
    if !is_cpu then
      cpu_timer := SOME (startCPUTimer ())
    else
      real_timer := SOME (startRealTimer ()));
  fun get_elapsed () = (
    if !is_cpu then
      #usr (!cpu_timer |> the |> checkCPUTimer)
    else
      (!real_timer |> the |> checkRealTimer));
  fun save_time s = (timings := ((s, get_elapsed ()) :: !timings));
  fun get_timings () = !timings;
end
›

paragraph ‹Optimized code printings›

definition
  "array_freeze' = array_freeze"

lemma [code]: "array_freeze a = array_freeze' a"
  unfolding array_freeze'_def ..

definition
  "array_unfreeze' = array_unfreeze"

lemma [code]: "array_unfreeze a = array_unfreeze' a"
  unfolding array_unfreeze'_def ..

definition
  "array_copy' = array_copy"

lemma [code]: "array_copy a = array_copy' a"
  unfolding array_copy'_def ..

code_printing constant array_freeze'  (SML) "(fn () => Array.vector _)"

code_printing constant array_unfreeze'  (SML)
  "(fn a => fn () => Array.tabulate (Vector.length a, fn i => Vector.sub (a, i))) _"

code_printing constant array_copy'  (SML)
  "(fn a => fn () => Array.tabulate (Array.length a, fn i => Array.sub (a, i))) _"

text ‹According to microbenchmarks, these versions are nearly two times faster.›

code_printing constant array_unfreeze'  (SML)
"(fn a => fn () =>
  if Vector.length a = 0
  then Array.fromList []
  else
    let
      val x = Vector.sub(a, 0)
      val n = Array.array (Vector.length a, x)
      val '_ = Array.copyVec {src=a,dst=n,di=0}
    in n end
) _"

code_printing constant array_copy'  (SML)
"(fn a => fn () =>
  if Array.length a = 0
  then Array.fromList []
  else
    let
      val x = Array.sub(a, 0)
      val n = Array.array (Array.length a, x)
      val '_ = Array.copy {src=a,dst=n,di=0}
    in n end
) _"


partial_function (heap) imp_for_int_inner ::
  "integer  integer  ('a  bool Heap)  (integer  'a  'a Heap)  'a  'a Heap" where
  "imp_for_int_inner i u c f s = (if i  u then return s else
    do {ctn <- c s; if ctn then f i s  imp_for_int_inner (i + 1) u c f else return s})"

lemma integer_of_nat_le_simp:
  "integer_of_nat i  integer_of_nat u  i  u"
  unfolding integer_of_nat_eq_of_nat by simp

lemma imp_for_imp_for_int_inner[code_unfold]:
  "imp_for i u c f s
  = imp_for_int_inner (integer_of_nat i) (integer_of_nat u) c (f o nat_of_integer) s"
  apply (induction "u - i" arbitrary: i u s)
   apply (simp add: integer_of_nat_le_simp imp_for_int_inner.simps; fail)
  apply (subst imp_for_int_inner.simps, simp add: integer_of_nat_le_simp)
  apply safe
  subgoal
    by auto
  apply (fo_rule arg_cong, rule ext)
  apply clarsimp
  apply (fo_rule arg_cong)
  apply (auto simp: algebra_simps integer_of_nat_eq_of_nat)
  done

definition
  "imp_for_int i u c f s 
  imp_for_int_inner (integer_of_nat i) (integer_of_nat u) c (f o nat_of_integer) s"

lemma imp_for_imp_for_int[code_unfold]:
  "imp_for = imp_for_int"
  by (intro ext, unfold imp_for_int_def imp_for_imp_for_int_inner, rule HOL.refl)

partial_function (heap) imp_for'_int_inner ::
  "integer  integer  (integer  'a  'a Heap)  'a  'a Heap" where
  "imp_for'_int_inner i u f s =
    (if i  u then return s else f i s  imp_for'_int_inner (i + 1) u f)"

lemma imp_for'_imp_for_int_inner:
  "imp_for' i u f s 
  imp_for'_int_inner (integer_of_nat i) (integer_of_nat u) (f o nat_of_integer) s"
  apply (induction "u - i" arbitrary: i u s)
   apply (simp add: integer_of_nat_le_simp imp_for'_int_inner.simps; fail)
  apply (subst imp_for'_int_inner.simps, simp add: integer_of_nat_le_simp)
  apply safe
  subgoal
    by auto
  apply (fo_rule arg_cong, rule ext)
  apply (auto simp: algebra_simps integer_of_nat_eq_of_nat)
  done

lemma imp_for'_int_cong:
  "imp_for' l u f a = imp_for' l u g a"
  if " i x. l  i  i < u  f i x = g i x"
  using that
  apply (induction "u - l" arbitrary: l u a)
   apply (simp add: imp_for'.simps; fail)
  apply (subst imp_for'.simps, simp add: integer_of_nat_le_simp)
  apply safe
  apply clarsimp
  apply (fo_rule arg_cong)
  apply auto
  done


definition
  "imp_for'_int i u f s 
  imp_for'_int_inner (integer_of_nat i) (integer_of_nat u) (f o nat_of_integer) s"

lemma imp_for'_imp_for_int[code_unfold, int_folds]:
  "imp_for' = imp_for'_int"
  by (intro ext, unfold imp_for'_int_def imp_for'_imp_for_int_inner, rule HOL.refl)

code_printing code_module "Iterators"  (SML)
‹
fun imp_for_inner i u c f s =
  let
    fun imp_for1 i u f s =
      if IntInf.<= (u, i) then (fn () => s)
      else if c s () then imp_for1 (i + 1) u f (f i s ())
      else (fn () => s)
  in imp_for1 i u f s end;

fun imp_fora_inner i u f s =
  let
    fun imp_for1 i u f s =
      if IntInf.<= (u, i) then (fn () => s)
      else imp_for1 (i + 1) u f (f i s ())
  in imp_for1 i u f s end;
›

code_reserved (SML) imp_for_inner imp_fora_inner
(* How much do we gain from this? *)
(*
code_printing
  constant imp_for_int_inner ⇀ (SML) "imp'_for'_inner"
| constant imp_for'_int_inner ⇀ (SML) "imp'_fora'_inner"
*)


definition
  "nth_integer a n = Array.nth a (nat_of_integer n)"

definition
  "upd_integer n x a = Array.upd (nat_of_integer n) x a"

code_printing
  constant upd_integer  (SML) "(fn a => fn () => (Array.update (a, IntInf.toInt _, _); a)) _"
| constant nth_integer  (SML) "(fn () => Array.sub (_, IntInf.toInt _))"

definition
  "fw_upd_impl_integer n a k i j = do {
  let n = n + 1;
  let i' = i * n + j;
  y  nth_integer a (i * n + k);
  z  nth_integer a (k * n + j);
  x  nth_integer a i';
  let m = dbm_add_int y z;
  if dbm_lt_int m x then upd_integer i' m a else return a
}
"

lemma nat_of_integer_add:
  "nat_of_integer i + nat_of_integer j = nat_of_integer (i + j)" if "i  0" "j  0"
proof -
  have *: "nat a + nat b = nat (a + b)" if "a  0" "b  0" for a b
    using that by simp
  show ?thesis
    unfolding nat_of_integer.rep_eq
    apply (simp add: *)
    apply (subst *)
    subgoal
      using zero_integer.rep_eq less_eq_integer.rep_eq that by metis
    subgoal
      using zero_integer.rep_eq less_eq_integer.rep_eq that by metis
    ..
qed

lemma nat_of_integer_mult:
  "nat_of_integer i * nat_of_integer j = nat_of_integer (i * j)" if "i  0" "j  0"
  using that
proof -
  have *: "nat a * nat b = nat (a * b)" if "a  0" "b  0" for a b
    using that by (simp add: nat_mult_distrib)
  show ?thesis
    unfolding nat_of_integer.rep_eq
    apply simp
    apply (subst *)
    subgoal
      using zero_integer.rep_eq less_eq_integer.rep_eq that by metis
    subgoal
      using zero_integer.rep_eq less_eq_integer.rep_eq that by metis
    ..
qed

lemma fw_upd_impl_int_fw_upd_impl_integer:
  "fw_upd_impl_int (nat_of_integer n) a (nat_of_integer k) (nat_of_integer i) (nat_of_integer j)
= fw_upd_impl_integer n a k i j"
  if "i  0" "j  0" "k  0" "n  0"
  unfolding
    fw_upd_impl_integer_def fw_upd_impl_int_def[symmetric] fw_upd_impl_def mtx_get_def mtx_set_def
  unfolding int_folds less
  unfolding nth_integer_def upd_integer_def fst_conv snd_conv
  using that
  apply (simp add: nat_of_integer_add nat_of_integer_mult algebra_simps)
  apply (fo_rule arg_cong | rule ext)+
  apply (simp add: nat_of_integer_add nat_of_integer_mult algebra_simps)
  done

lemma integer_of_nat_nat_of_integer:
  "integer_of_nat (nat_of_integer n) = n" if "n  0"
  using that by (simp add: integer_of_nat_eq_of_nat max_def)

lemma integer_of_nat_aux:
  "integer_of_nat (nat_of_integer n + 1) = n + 1" if "n  0"
proof -
  have "nat_of_integer n + 1 = nat_of_integer n + nat_of_integer 1"
    by simp
  also have " = nat_of_integer (n + 1)"
    using that by (subst nat_of_integer_add; simp)
  finally show ?thesis
    using that by (simp add: integer_of_nat_nat_of_integer)
qed

definition
  "fwi_impl_integer n a k = fwi_impl_int (nat_of_integer n) a (nat_of_integer k)"

lemma fwi_impl_int_eq1:
  "fwi_impl_int n a k  fwi_impl_integer (integer_of_nat n) a (integer_of_nat k)"
  unfolding fwi_impl_integer_def fwi_impl_int_def by simp

partial_function (heap) imp_for'_int_int ::
  "int  int  (int  'a  'a Heap)  'a  'a Heap" where
  "imp_for'_int_int i u f s =
    (if i  u then return s else f i s  imp_for'_int_int (i + 1) u f)"

lemma int_bounds_up_induct:
  assumes "l. u  (l :: int)  P l u"
      and "l. P (l + 1) u  l < u  P l u"
  shows "P l u"
  by (induction "nat (u - l)" arbitrary: l) (auto intro: assms)

lemma imp_for'_int_int_cong:
  "imp_for'_int_int l u f a = imp_for'_int_int l u g a"
  if " i x. l  i  i < u  f i x = g i x"
  using that
  apply (induction arbitrary: a rule: int_bounds_up_induct)
   apply (simp add: imp_for'_int_int.simps; fail)
  apply (subst (2) imp_for'_int_int.simps)
  apply (subst imp_for'_int_int.simps)
  apply simp
  apply (fo_rule arg_cong, rule ext)
  .

lemma plus_int_int_of_integer_aux:
  "(plus_int (int_of_integer l) 1) = int_of_integer (l + 1)"
  by simp

lemma imp_for'_int_inner_imp_for'_int_int:
  "imp_for'_int_inner l u f a
= imp_for'_int_int (int_of_integer l) (int_of_integer u) (f o integer_of_int) a"
  apply (induction "int_of_integer l" "int_of_integer u" arbitrary: a l u rule: int_bounds_up_induct)
   apply (simp add: imp_for'_int_int.simps imp_for'_int_inner.simps less_eq_integer.rep_eq; fail)
  apply (subst imp_for'_int_int.simps)
  apply (subst imp_for'_int_inner.simps)
  apply (simp add: less_eq_integer.rep_eq)
  apply (fo_rule arg_cong, rule ext)
  apply (subst plus_int_int_of_integer_aux)
  apply rprems
  using plus_int_int_of_integer_aux by metis+

lemma imp_for'_int_inner_cong:
  "imp_for'_int_inner l u f a = imp_for'_int_inner l u g a"
  if " i x. l  i  i < u  f i x = g i x"
  unfolding imp_for'_int_inner_imp_for'_int_int
  by (rule imp_for'_int_int_cong)(auto simp: less_eq_integer.rep_eq less_integer.rep_eq intro: that)

schematic_goal fwi_impl_int_unfold1:
  "fwi_impl_int (nat_of_integer n) a (nat_of_integer k) = ?i" if "0  k" "0  n"
  unfolding fwi_impl_int_def[symmetric] fwi_impl_def
  unfolding int_folds
  unfolding imp_for'_int_def
  unfolding comp_def integer_of_nat_0
  apply (rule imp_for'_int_inner_cong)
  apply (rule imp_for'_int_inner_cong)
  apply (rule fw_upd_impl_int_fw_upd_impl_integer)
  by (assumption | rule that)+

lemma integer_of_nat_add:
  "integer_of_nat (x + y) = integer_of_nat x + integer_of_nat y"
  by (simp add: integer_of_nat_eq_of_nat)

schematic_goal fwi_impl_int_code [code]:
  "fwi_impl_int n a k  ?x"
  unfolding fwi_impl_int_eq1
  unfolding fwi_impl_integer_def
  apply (subst fwi_impl_int_unfold1)
    apply (simp add: integer_of_nat_eq_of_nat; fail)
   apply (simp add: integer_of_nat_eq_of_nat; fail)
  unfolding nat_of_integer_integer_of_nat
  unfolding integer_of_nat_add integer_of_nat_1
  apply (abstract_let "integer_of_nat n + 1" n_plus_1)
  apply (abstract_let "integer_of_nat n" n)
  .


code_printing code_module "Integer"  (Eval)
‹
structure Integer: sig
  val div_mod: int -> int -> int * int
end = struct
  fun div_mod i j = (i div j, i mod j)
end
›

text ‹Delete unused code module.›
code_printing code_module Bits_Integer  (SML) ‹›

text ‹For agreement with SML›
code_printing
  type_constructor Typerep.typerep  (Eval)
| constant Typerep.Typerep  (Eval)

(*
code_printing
  type_constructor Typerep.typerep ⇀ (Eval) "typerepa"
| constant Typerep.Typerep ⇀ (Eval) "Typerep/ (_, _)"
*)

(* datatype typerepa = Typerep of string * typerepa list; *)

(* Speedup for Poly *)
(*
definition
  "fw_upd_impl_integer' = fw_upd_impl_integer"

lemma [code]:
  "fw_upd_impl_integer = fw_upd_impl_integer'"
  unfolding fw_upd_impl_integer'_def ..

code_printing
  constant fw_upd_impl_integer' ⇀ (SML) "(fn n => fn a => fn k => fn i => fn j =>
  let
    val na = IntInf.+ (n, (1 : IntInf.int));
    val ia = IntInf.+ (IntInf.* (i, na), j);
    val x = Array.sub (a, IntInf.toInt ia);
    val y = Array.sub (a, IntInf.toInt (IntInf.+ (IntInf.* (i, na), k)));
    val z = Array.sub (a, IntInf.toInt (IntInf.+ (IntInf.* (k, na), j)));
    val m = dbm'_add'_int y z;
  in
    if dbm'_lt'_int m x
    then (fn () => (Array.update (a, IntInf.toInt ia, m); a))
    else (fn () => a)
  end) _ _ _ _ _"
*)

lemmas [code] = imp_for'_int_inner.simps imp_for_int_inner.simps

text ‹We eliminate this module because it uses constants that are only implemented by IntInf›
and not Int›. When compiling we will use a hack to change IntInf› with Int› (for efficiency)
and therefore this module would break the hack.›
code_printing code_module Bit_Shifts  (SML) ‹›

text ‹Using the Eval› code target instead of SML› makes it easier to replace IntInf› with Int›.›
export_code parse_convert_check parse_convert_run_print parse_convert_run_check Result Error
  nat_of_integer int_of_integer DBMEntry.Le DBMEntry.Lt DBMEntry.INF
  Impl1 Impl2 Impl3 Buechi Debug Reachable_Set Buechi_Set
  E_op_impl
  in Eval module_name Model_Checker file_prefix "Certificate"

export_code parse_convert_check parse_convert_run_print parse_convert_run_check Result Error
  nat_of_integer int_of_integer DBMEntry.Le DBMEntry.Lt DBMEntry.INF
  Impl1 Impl2 Impl3 Buechi Reachable_Set Buechi_Set
  E_op_impl
  in SML module_name Model_Checker



code_printing code_module "Printing"  (Haskell)
‹
import qualified Debug.Trace;

print s = Debug.Trace.trace s ();

printM s = Debug.Trace.traceM s;
›

code_printing
  constant Printing.print  (Haskell) "Printing.print _"

code_printing
  constant print_line_impl  (Haskell) "Printing.printM _"

(* Timing Utilities For Haskell *)
(*
code_printing code_module "Timing" ⇀ (Haskell)
‹
import Data.Time.Clock.System;

now = systemToTAITime . Prelude.unsafePerformIO getSystemTime;
›
*)

(* code_printing
  constant "now" ⇀ (Haskell) "Prelude.const (Time (Int'_of'_integer 0)) _"

code_printing
  constant "time_to_string" ⇀ (Haskell) "Prelude.show _"

code_printing
  constant "(-) :: time ⇒ time ⇒ time" ⇀ (Haskell)
    "(case _ of Time a -> case _ of Time b -> Time (b - a))" *)

code_printing
  type_constructor time  (Haskell) "Integer"
  | constant "now"  (Haskell) "Prelude.const 0"
  | constant "time_to_string"  (Haskell) "Prelude.show _"
  | constant "(-) :: time  time  time"  (Haskell) "(-)"

code_printing
  constant list_of_set'  (Haskell) "(case _ of Set xs -> xs)"

export_code parse_convert_check parse_convert_run_print parse_convert_run_check Result Error
  nat_of_integer int_of_integer DBMEntry.Le DBMEntry.Lt DBMEntry.INF
  Impl1 Impl2 Impl3 Buechi Reachable_Set Buechi_Set
  in Haskell module_name Model_Checker (*file "../Haskell/"*)

end