CFEM.bandmat

From Coq Require Import String List ZArith.
From compcert Require Import Coqlib Integers Floats AST Ctypes Cop Clight Clightdefs.
Import Clightdefs.ClightNotations.
Local Open Scope Z_scope.
Local Open Scope string_scope.
Local Open Scope clight_scope.

Module Info.
  Definition version := "3.15".
  Definition build_number := "".
  Definition build_tag := "".
  Definition build_branch := "".
  Definition arch := "aarch64".
  Definition model := "default".
  Definition abi := "apple".
  Definition bitsize := 64.
  Definition big_endian := false.
  Definition source_file := "../src/bandmat.c".
  Definition normalized := true.
End Info.

Definition _A : ident := $"A".
Definition _P : ident := $"P".
Definition _PA : ident := $"PA".
Definition _PR : ident := $"PR".
Definition ___builtin_annot : ident := $"__builtin_annot".
Definition ___builtin_annot_intval : ident := $"__builtin_annot_intval".
Definition ___builtin_bswap : ident := $"__builtin_bswap".
Definition ___builtin_bswap16 : ident := $"__builtin_bswap16".
Definition ___builtin_bswap32 : ident := $"__builtin_bswap32".
Definition ___builtin_bswap64 : ident := $"__builtin_bswap64".
Definition ___builtin_cls : ident := $"__builtin_cls".
Definition ___builtin_clsl : ident := $"__builtin_clsl".
Definition ___builtin_clsll : ident := $"__builtin_clsll".
Definition ___builtin_clz : ident := $"__builtin_clz".
Definition ___builtin_clzl : ident := $"__builtin_clzl".
Definition ___builtin_clzll : ident := $"__builtin_clzll".
Definition ___builtin_ctz : ident := $"__builtin_ctz".
Definition ___builtin_ctzl : ident := $"__builtin_ctzl".
Definition ___builtin_ctzll : ident := $"__builtin_ctzll".
Definition ___builtin_debug : ident := $"__builtin_debug".
Definition ___builtin_expect : ident := $"__builtin_expect".
Definition ___builtin_fabs : ident := $"__builtin_fabs".
Definition ___builtin_fabsf : ident := $"__builtin_fabsf".
Definition ___builtin_fmadd : ident := $"__builtin_fmadd".
Definition ___builtin_fmax : ident := $"__builtin_fmax".
Definition ___builtin_fmin : ident := $"__builtin_fmin".
Definition ___builtin_fmsub : ident := $"__builtin_fmsub".
Definition ___builtin_fnmadd : ident := $"__builtin_fnmadd".
Definition ___builtin_fnmsub : ident := $"__builtin_fnmsub".
Definition ___builtin_fsqrt : ident := $"__builtin_fsqrt".
Definition ___builtin_membar : ident := $"__builtin_membar".
Definition ___builtin_memcpy_aligned : ident := $"__builtin_memcpy_aligned".
Definition ___builtin_sel : ident := $"__builtin_sel".
Definition ___builtin_sqrt : ident := $"__builtin_sqrt".
Definition ___builtin_unreachable : ident := $"__builtin_unreachable".
Definition ___builtin_va_arg : ident := $"__builtin_va_arg".
Definition ___builtin_va_copy : ident := $"__builtin_va_copy".
Definition ___builtin_va_end : ident := $"__builtin_va_end".
Definition ___builtin_va_start : ident := $"__builtin_va_start".
Definition ___compcert_i64_dtos : ident := $"__compcert_i64_dtos".
Definition ___compcert_i64_dtou : ident := $"__compcert_i64_dtou".
Definition ___compcert_i64_sar : ident := $"__compcert_i64_sar".
Definition ___compcert_i64_sdiv : ident := $"__compcert_i64_sdiv".
Definition ___compcert_i64_shl : ident := $"__compcert_i64_shl".
Definition ___compcert_i64_shr : ident := $"__compcert_i64_shr".
Definition ___compcert_i64_smod : ident := $"__compcert_i64_smod".
Definition ___compcert_i64_smulh : ident := $"__compcert_i64_smulh".
Definition ___compcert_i64_stod : ident := $"__compcert_i64_stod".
Definition ___compcert_i64_stof : ident := $"__compcert_i64_stof".
Definition ___compcert_i64_udiv : ident := $"__compcert_i64_udiv".
Definition ___compcert_i64_umod : ident := $"__compcert_i64_umod".
Definition ___compcert_i64_umulh : ident := $"__compcert_i64_umulh".
Definition ___compcert_i64_utod : ident := $"__compcert_i64_utod".
Definition ___compcert_i64_utof : ident := $"__compcert_i64_utof".
Definition ___compcert_va_composite : ident := $"__compcert_va_composite".
Definition ___compcert_va_float64 : ident := $"__compcert_va_float64".
Definition ___compcert_va_int32 : ident := $"__compcert_va_int32".
Definition ___compcert_va_int64 : ident := $"__compcert_va_int64".
Definition ___stringlit_1 : ident := $"__stringlit_1".
Definition ___stringlit_2 : ident := $"__stringlit_2".
Definition ___stringlit_3 : ident := $"__stringlit_3".
Definition ___stringlit_4 : ident := $"__stringlit_4".
Definition ___stringlit_5 : ident := $"__stringlit_5".
Definition _abort : ident := $"abort".
Definition _b : ident := $"b".
Definition _bandmat_addto : ident := $"bandmat_addto".
Definition _bandmat_clear : ident := $"bandmat_clear".
Definition _bandmat_factor : ident := $"bandmat_factor".
Definition _bandmat_free : ident := $"bandmat_free".
Definition _bandmat_get : ident := $"bandmat_get".
Definition _bandmat_malloc : ident := $"bandmat_malloc".
Definition _bandmat_norm : ident := $"bandmat_norm".
Definition _bandmat_norm2 : ident := $"bandmat_norm2".
Definition _bandmat_print : ident := $"bandmat_print".
Definition _bandmat_set : ident := $"bandmat_set".
Definition _bandmat_solve : ident := $"bandmat_solve".
Definition _bandmat_t : ident := $"bandmat_t".
Definition _bandmatn_addto : ident := $"bandmatn_addto".
Definition _bandmatn_clear : ident := $"bandmatn_clear".
Definition _bandmatn_get : ident := $"bandmatn_get".
Definition _bandmatn_set : ident := $"bandmatn_set".
Definition _bi : ident := $"bi".
Definition _bw : ident := $"bw".
Definition _d : ident := $"d".
Definition _data : ident := $"data".
Definition _data_norm : ident := $"data_norm".
Definition _data_norm2 : ident := $"data_norm2".
Definition _dense_to_band : ident := $"dense_to_band".
Definition _densemat_get : ident := $"densemat_get".
Definition _densemat_t : ident := $"densemat_t".
Definition _dj : ident := $"dj".
Definition _dm : ident := $"dm".
Definition _free : ident := $"free".
Definition _i : ident := $"i".
Definition _i__1 : ident := $"i__1".
Definition _j : ident := $"j".
Definition _j__1 : ident := $"j__1".
Definition _k : ident := $"k".
Definition _m : ident := $"m".
Definition _main : ident := $"main".
Definition _memset : ident := $"memset".
Definition _n : ident := $"n".
Definition _printf : ident := $"printf".
Definition _rows : ident := $"rows".
Definition _sqrt : ident := $"sqrt".
Definition _surely_malloc : ident := $"surely_malloc".
Definition _vm : ident := $"vm".
Definition _x : ident := $"x".
Definition _yi : ident := $"yi".
Definition _t'1 : ident := 128%positive.
Definition _t'2 : ident := 129%positive.
Definition _t'3 : ident := 130%positive.
Definition _t'4 : ident := 131%positive.
Definition _t'5 : ident := 132%positive.
Definition _t'6 : ident := 133%positive.
Definition _t'7 : ident := 134%positive.
Definition _t'8 : ident := 135%positive.
Definition _t'9 : ident := 136%positive.

Definition v___stringlit_3 := {|
  gvar_info := (tarray tschar 25);
  gvar_init := (Init_int8 (Int.repr 98) :: Init_int8 (Int.repr 97) ::
                Init_int8 (Int.repr 110) :: Init_int8 (Int.repr 100) ::
                Init_int8 (Int.repr 109) :: Init_int8 (Int.repr 97) ::
                Init_int8 (Int.repr 116) :: Init_int8 (Int.repr 95) ::
                Init_int8 (Int.repr 103) :: Init_int8 (Int.repr 101) ::
                Init_int8 (Int.repr 116) :: Init_int8 (Int.repr 40) ::
                Init_int8 (Int.repr 80) :: Init_int8 (Int.repr 65) ::
                Init_int8 (Int.repr 44) :: Init_int8 (Int.repr 107) ::
                Init_int8 (Int.repr 44) :: Init_int8 (Int.repr 48) ::
                Init_int8 (Int.repr 41) :: Init_int8 (Int.repr 32) ::
                Init_int8 (Int.repr 62) :: Init_int8 (Int.repr 61) ::
                Init_int8 (Int.repr 32) :: Init_int8 (Int.repr 48) ::
                Init_int8 (Int.repr 0) :: nil);
  gvar_readonly := true;
  gvar_volatile := false
|}.

Definition v___stringlit_1 := {|
  gvar_info := (tarray tschar 9);
  gvar_init := (Init_int8 (Int.repr 32) :: Init_int8 (Int.repr 32) ::
                Init_int8 (Int.repr 37) :: Init_int8 (Int.repr 32) ::
                Init_int8 (Int.repr 54) :: Init_int8 (Int.repr 46) ::
                Init_int8 (Int.repr 51) :: Init_int8 (Int.repr 103) ::
                Init_int8 (Int.repr 0) :: nil);
  gvar_readonly := true;
  gvar_volatile := false
|}.

Definition v___stringlit_5 := {|
  gvar_info := (tarray tschar 30);
  gvar_init := (Init_int8 (Int.repr 37) :: Init_int8 (Int.repr 115) ::
                Init_int8 (Int.repr 58) :: Init_int8 (Int.repr 37) ::
                Init_int8 (Int.repr 100) :: Init_int8 (Int.repr 58) ::
                Init_int8 (Int.repr 32) :: Init_int8 (Int.repr 102) ::
                Init_int8 (Int.repr 97) :: Init_int8 (Int.repr 105) ::
                Init_int8 (Int.repr 108) :: Init_int8 (Int.repr 101) ::
                Init_int8 (Int.repr 100) :: Init_int8 (Int.repr 32) ::
                Init_int8 (Int.repr 97) :: Init_int8 (Int.repr 115) ::
                Init_int8 (Int.repr 115) :: Init_int8 (Int.repr 101) ::
                Init_int8 (Int.repr 114) :: Init_int8 (Int.repr 116) ::
                Init_int8 (Int.repr 105) :: Init_int8 (Int.repr 111) ::
                Init_int8 (Int.repr 110) :: Init_int8 (Int.repr 32) ::
                Init_int8 (Int.repr 96) :: Init_int8 (Int.repr 37) ::
                Init_int8 (Int.repr 115) :: Init_int8 (Int.repr 39) ::
                Init_int8 (Int.repr 10) :: Init_int8 (Int.repr 0) :: nil);
  gvar_readonly := true;
  gvar_volatile := false
|}.

Definition v___stringlit_2 := {|
  gvar_info := (tarray tschar 2);
  gvar_init := (Init_int8 (Int.repr 10) :: Init_int8 (Int.repr 0) :: nil);
  gvar_readonly := true;
  gvar_volatile := false
|}.

Definition v___stringlit_4 := {|
  gvar_info := (tarray tschar 10);
  gvar_init := (Init_int8 (Int.repr 98) :: Init_int8 (Int.repr 97) ::
                Init_int8 (Int.repr 110) :: Init_int8 (Int.repr 100) ::
                Init_int8 (Int.repr 109) :: Init_int8 (Int.repr 97) ::
                Init_int8 (Int.repr 116) :: Init_int8 (Int.repr 46) ::
                Init_int8 (Int.repr 99) :: Init_int8 (Int.repr 0) :: nil);
  gvar_readonly := true;
  gvar_volatile := false
|}.

Definition f_bandmat_malloc := {|
  fn_return := (tptr (Tstruct _bandmat_t noattr));
  fn_callconv := cc_default;
  fn_params := ((_n, tint) :: (_b, tint) :: nil);
  fn_vars := nil;
  fn_temps := ((_vm, (tptr (Tstruct _bandmat_t noattr))) ::
               (_t'1, (tptr tvoid)) :: nil);
  fn_body :=
(Ssequence
  (Ssequence
    (Scall (Some _t'1)
      (Evar _surely_malloc (Tfunction (tulong :: nil) (tptr tvoid)
                             cc_default))
      ((Ebinop Oadd (Esizeof (Tstruct _bandmat_t noattr) tulong)
         (Ebinop Omul
           (Ebinop Omul (Etempvar _n tint)
             (Ebinop Oadd (Etempvar _b tint) (Econst_int (Int.repr 1) tint)
               tint) tint) (Esizeof tdouble tulong) tulong) tulong) :: nil))
    (Sset _vm (Etempvar _t'1 (tptr tvoid))))
  (Ssequence
    (Sassign
      (Efield
        (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
          (Tstruct _bandmat_t noattr)) _m tint) (Etempvar _n tint))
    (Ssequence
      (Sassign
        (Efield
          (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
            (Tstruct _bandmat_t noattr)) _b tint) (Etempvar _b tint))
      (Sreturn (Some (Etempvar _vm (tptr (Tstruct _bandmat_t noattr))))))))
|}.

Definition f_bandmat_free := {|
  fn_return := tvoid;
  fn_callconv := cc_default;
  fn_params := ((_vm, (tptr (Tstruct _bandmat_t noattr))) :: nil);
  fn_vars := nil;
  fn_temps := nil;
  fn_body :=
(Scall None (Evar _free (Tfunction ((tptr tvoid) :: nil) tvoid cc_default))
  ((Etempvar _vm (tptr (Tstruct _bandmat_t noattr))) :: nil))
|}.

Definition f_bandmatn_clear := {|
  fn_return := tvoid;
  fn_callconv := cc_default;
  fn_params := ((_data, (tptr tdouble)) :: (_m, tint) :: (_b, tint) :: nil);
  fn_vars := nil;
  fn_temps := nil;
  fn_body :=
(Scall None
  (Evar _memset (Tfunction ((tptr tvoid) :: tint :: tulong :: nil)
                  (tptr tvoid) cc_default))
  ((Etempvar _data (tptr tdouble)) :: (Econst_int (Int.repr 0) tint) ::
   (Ebinop Omul
     (Ebinop Omul (Etempvar _m tint)
       (Ebinop Oadd (Etempvar _b tint) (Econst_int (Int.repr 1) tint) tint)
       tint) (Esizeof tdouble tulong) tulong) :: nil))
|}.

Definition f_bandmat_clear := {|
  fn_return := tvoid;
  fn_callconv := cc_default;
  fn_params := ((_vm, (tptr (Tstruct _bandmat_t noattr))) :: nil);
  fn_vars := nil;
  fn_temps := ((_t'2, tint) :: (_t'1, tint) :: nil);
  fn_body :=
(Ssequence
  (Sset _t'1
    (Efield
      (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
        (Tstruct _bandmat_t noattr)) _m tint))
  (Ssequence
    (Sset _t'2
      (Efield
        (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
          (Tstruct _bandmat_t noattr)) _b tint))
    (Scall None
      (Evar _bandmatn_clear (Tfunction
                              ((tptr tdouble) :: tint :: tint :: nil) tvoid
                              cc_default))
      ((Efield
         (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
           (Tstruct _bandmat_t noattr)) _data (tarray tdouble 0)) ::
       (Etempvar _t'1 tint) :: (Etempvar _t'2 tint) :: nil))))
|}.

Definition f_bandmatn_get := {|
  fn_return := tdouble;
  fn_callconv := cc_default;
  fn_params := ((_data, (tptr tdouble)) :: (_rows, tint) :: (_i, tint) ::
                (_d, tint) :: nil);
  fn_vars := nil;
  fn_temps := ((_t'1, tdouble) :: nil);
  fn_body :=
(Ssequence
  (Sset _t'1
    (Ederef
      (Ebinop Oadd (Etempvar _data (tptr tdouble))
        (Ebinop Oadd (Etempvar _i tint)
          (Ebinop Omul (Etempvar _d tint) (Etempvar _rows tint) tint) tint)
        (tptr tdouble)) tdouble))
  (Sreturn (Some (Etempvar _t'1 tdouble))))
|}.

Definition f_bandmatn_set := {|
  fn_return := tvoid;
  fn_callconv := cc_default;
  fn_params := ((_data, (tptr tdouble)) :: (_rows, tint) :: (_i, tint) ::
                (_d, tint) :: (_x, tdouble) :: nil);
  fn_vars := nil;
  fn_temps := nil;
  fn_body :=
(Sassign
  (Ederef
    (Ebinop Oadd (Etempvar _data (tptr tdouble))
      (Ebinop Oadd (Etempvar _i tint)
        (Ebinop Omul (Etempvar _d tint) (Etempvar _rows tint) tint) tint)
      (tptr tdouble)) tdouble) (Etempvar _x tdouble))
|}.

Definition f_bandmatn_addto := {|
  fn_return := tvoid;
  fn_callconv := cc_default;
  fn_params := ((_data, (tptr tdouble)) :: (_rows, tint) :: (_i, tint) ::
                (_d, tint) :: (_x, tdouble) :: nil);
  fn_vars := nil;
  fn_temps := ((_t'1, tdouble) :: nil);
  fn_body :=
(Ssequence
  (Sset _t'1
    (Ederef
      (Ebinop Oadd (Etempvar _data (tptr tdouble))
        (Ebinop Oadd (Etempvar _i tint)
          (Ebinop Omul (Etempvar _d tint) (Etempvar _rows tint) tint) tint)
        (tptr tdouble)) tdouble))
  (Sassign
    (Ederef
      (Ebinop Oadd (Etempvar _data (tptr tdouble))
        (Ebinop Oadd (Etempvar _i tint)
          (Ebinop Omul (Etempvar _d tint) (Etempvar _rows tint) tint) tint)
        (tptr tdouble)) tdouble)
    (Ebinop Oadd (Etempvar _t'1 tdouble) (Etempvar _x tdouble) tdouble)))
|}.

Definition f_bandmat_get := {|
  fn_return := tdouble;
  fn_callconv := cc_default;
  fn_params := ((_dm, (tptr (Tstruct _bandmat_t noattr))) :: (_i, tint) ::
                (_d, tint) :: nil);
  fn_vars := nil;
  fn_temps := ((_t'1, tdouble) :: (_t'2, tint) :: nil);
  fn_body :=
(Ssequence
  (Ssequence
    (Sset _t'2
      (Efield
        (Ederef (Etempvar _dm (tptr (Tstruct _bandmat_t noattr)))
          (Tstruct _bandmat_t noattr)) _m tint))
    (Scall (Some _t'1)
      (Evar _bandmatn_get (Tfunction
                            ((tptr tdouble) :: tint :: tint :: tint :: nil)
                            tdouble cc_default))
      ((Efield
         (Ederef (Etempvar _dm (tptr (Tstruct _bandmat_t noattr)))
           (Tstruct _bandmat_t noattr)) _data (tarray tdouble 0)) ::
       (Etempvar _t'2 tint) :: (Etempvar _i tint) :: (Etempvar _d tint) ::
       nil)))
  (Sreturn (Some (Etempvar _t'1 tdouble))))
|}.

Definition f_bandmat_set := {|
  fn_return := tvoid;
  fn_callconv := cc_default;
  fn_params := ((_dm, (tptr (Tstruct _bandmat_t noattr))) :: (_i, tint) ::
                (_d, tint) :: (_x, tdouble) :: nil);
  fn_vars := nil;
  fn_temps := ((_t'1, tint) :: nil);
  fn_body :=
(Ssequence
  (Sset _t'1
    (Efield
      (Ederef (Etempvar _dm (tptr (Tstruct _bandmat_t noattr)))
        (Tstruct _bandmat_t noattr)) _m tint))
  (Scall None
    (Evar _bandmatn_set (Tfunction
                          ((tptr tdouble) :: tint :: tint :: tint ::
                           tdouble :: nil) tvoid cc_default))
    ((Efield
       (Ederef (Etempvar _dm (tptr (Tstruct _bandmat_t noattr)))
         (Tstruct _bandmat_t noattr)) _data (tarray tdouble 0)) ::
     (Etempvar _t'1 tint) :: (Etempvar _i tint) :: (Etempvar _d tint) ::
     (Etempvar _x tdouble) :: nil)))
|}.

Definition f_bandmat_addto := {|
  fn_return := tvoid;
  fn_callconv := cc_default;
  fn_params := ((_dm, (tptr (Tstruct _bandmat_t noattr))) :: (_i, tint) ::
                (_d, tint) :: (_x, tdouble) :: nil);
  fn_vars := nil;
  fn_temps := ((_t'1, tint) :: nil);
  fn_body :=
(Ssequence
  (Sset _t'1
    (Efield
      (Ederef (Etempvar _dm (tptr (Tstruct _bandmat_t noattr)))
        (Tstruct _bandmat_t noattr)) _m tint))
  (Scall None
    (Evar _bandmatn_addto (Tfunction
                            ((tptr tdouble) :: tint :: tint :: tint ::
                             tdouble :: nil) tvoid cc_default))
    ((Efield
       (Ederef (Etempvar _dm (tptr (Tstruct _bandmat_t noattr)))
         (Tstruct _bandmat_t noattr)) _data (tarray tdouble 0)) ::
     (Etempvar _t'1 tint) :: (Etempvar _i tint) :: (Etempvar _d tint) ::
     (Etempvar _x tdouble) :: nil)))
|}.

Definition f_dense_to_band := {|
  fn_return := (tptr (Tstruct _bandmat_t noattr));
  fn_callconv := cc_default;
  fn_params := ((_A, (tptr (Tstruct _densemat_t noattr))) :: (_bw, tint) ::
                nil);
  fn_vars := nil;
  fn_temps := ((_n, tint) :: (_P, (tptr (Tstruct _bandmat_t noattr))) ::
               (_d, tint) :: (_j, tint) :: (_i, tint) :: (_t'2, tdouble) ::
               (_t'1, (tptr (Tstruct _bandmat_t noattr))) :: nil);
  fn_body :=
(Ssequence
  (Sset _n
    (Efield
      (Ederef (Etempvar _A (tptr (Tstruct _densemat_t noattr)))
        (Tstruct _densemat_t noattr)) _n tint))
  (Ssequence
    (Ssequence
      (Scall (Some _t'1)
        (Evar _bandmat_malloc (Tfunction (tint :: tint :: nil)
                                (tptr (Tstruct _bandmat_t noattr))
                                cc_default))
        ((Etempvar _n tint) :: (Etempvar _bw tint) :: nil))
      (Sset _P (Etempvar _t'1 (tptr (Tstruct _bandmat_t noattr)))))
    (Ssequence
      (Ssequence
        (Sset _d (Econst_int (Int.repr 0) tint))
        (Sloop
          (Ssequence
            (Sifthenelse (Ebinop Ole (Etempvar _d tint) (Etempvar _bw tint)
                           tint)
              Sskip
              Sbreak)
            (Ssequence
              (Sset _j (Etempvar _d tint))
              (Sloop
                (Ssequence
                  (Sifthenelse (Ebinop Olt (Etempvar _j tint)
                                 (Etempvar _n tint) tint)
                    Sskip
                    Sbreak)
                  (Ssequence
                    (Sset _i
                      (Ebinop Osub (Etempvar _j tint) (Etempvar _d tint)
                        tint))
                    (Ssequence
                      (Scall (Some _t'2)
                        (Evar _densemat_get (Tfunction
                                              ((tptr (Tstruct _densemat_t noattr)) ::
                                               tint :: tint :: nil) tdouble
                                              cc_default))
                        ((Etempvar _A (tptr (Tstruct _densemat_t noattr))) ::
                         (Etempvar _i tint) :: (Etempvar _j tint) :: nil))
                      (Scall None
                        (Evar _bandmat_set (Tfunction
                                             ((tptr (Tstruct _bandmat_t noattr)) ::
                                              tint :: tint :: tdouble :: nil)
                                             tvoid cc_default))
                        ((Etempvar _P (tptr (Tstruct _bandmat_t noattr))) ::
                         (Etempvar _j tint) :: (Etempvar _d tint) ::
                         (Etempvar _t'2 tdouble) :: nil)))))
                (Sset _j
                  (Ebinop Oadd (Etempvar _j tint)
                    (Econst_int (Int.repr 1) tint) tint)))))
          (Sset _d
            (Ebinop Oadd (Etempvar _d tint) (Econst_int (Int.repr 1) tint)
              tint))))
      (Sreturn (Some (Etempvar _P (tptr (Tstruct _bandmat_t noattr))))))))
|}.

Definition f_bandmat_print := {|
  fn_return := tvoid;
  fn_callconv := cc_default;
  fn_params := ((_PA, (tptr (Tstruct _bandmat_t noattr))) :: nil);
  fn_vars := nil;
  fn_temps := ((_n, tint) :: (_bw, tint) :: (_i, tint) :: (_d, tint) ::
               (_t'2, tdouble) :: (_t'1, tint) :: nil);
  fn_body :=
(Ssequence
  (Sset _n
    (Efield
      (Ederef (Etempvar _PA (tptr (Tstruct _bandmat_t noattr)))
        (Tstruct _bandmat_t noattr)) _m tint))
  (Ssequence
    (Sset _bw
      (Efield
        (Ederef (Etempvar _PA (tptr (Tstruct _bandmat_t noattr)))
          (Tstruct _bandmat_t noattr)) _b tint))
    (Ssequence
      (Sset _i (Econst_int (Int.repr 0) tint))
      (Sloop
        (Ssequence
          (Sifthenelse (Ebinop Olt (Etempvar _i tint) (Etempvar _n tint)
                         tint)
            Sskip
            Sbreak)
          (Ssequence
            (Ssequence
              (Sset _d (Econst_int (Int.repr 0) tint))
              (Sloop
                (Ssequence
                  (Ssequence
                    (Sifthenelse (Ebinop Ole (Etempvar _d tint)
                                   (Etempvar _bw tint) tint)
                      (Sset _t'1
                        (Ecast
                          (Ebinop Ole (Etempvar _d tint) (Etempvar _i tint)
                            tint) tbool))
                      (Sset _t'1 (Econst_int (Int.repr 0) tint)))
                    (Sifthenelse (Etempvar _t'1 tint) Sskip Sbreak))
                  (Ssequence
                    (Scall (Some _t'2)
                      (Evar _bandmat_get (Tfunction
                                           ((tptr (Tstruct _bandmat_t noattr)) ::
                                            tint :: tint :: nil) tdouble
                                           cc_default))
                      ((Etempvar _PA (tptr (Tstruct _bandmat_t noattr))) ::
                       (Etempvar _i tint) :: (Etempvar _d tint) :: nil))
                    (Scall None
                      (Evar _printf (Tfunction ((tptr tschar) :: nil) tint
                                      {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|}))
                      ((Evar ___stringlit_1 (tarray tschar 9)) ::
                       (Etempvar _t'2 tdouble) :: nil))))
                (Sset _d
                  (Ebinop Oadd (Etempvar _d tint)
                    (Econst_int (Int.repr 1) tint) tint))))
            (Scall None
              (Evar _printf (Tfunction ((tptr tschar) :: nil) tint
                              {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|}))
              ((Evar ___stringlit_2 (tarray tschar 2)) :: nil))))
        (Sset _i
          (Ebinop Oadd (Etempvar _i tint) (Econst_int (Int.repr 1) tint)
            tint))))))
|}.

Definition f_bandmat_factor := {|
  fn_return := tvoid;
  fn_callconv := cc_default;
  fn_params := ((_PA, (tptr (Tstruct _bandmat_t noattr))) :: nil);
  fn_vars := nil;
  fn_temps := ((_n, tint) :: (_bw, tint) :: (_k, tint) :: (_j, tint) ::
               (_j__1, tint) :: (_i, tint) :: (_t'9, tdouble) ::
               (_t'8, tdouble) :: (_t'7, tint) :: (_t'6, tdouble) ::
               (_t'5, tdouble) :: (_t'4, tint) :: (_t'3, tdouble) ::
               (_t'2, tdouble) :: (_t'1, tdouble) :: nil);
  fn_body :=
(Ssequence
  (Sset _n
    (Efield
      (Ederef (Etempvar _PA (tptr (Tstruct _bandmat_t noattr)))
        (Tstruct _bandmat_t noattr)) _m tint))
  (Ssequence
    (Sset _bw
      (Efield
        (Ederef (Etempvar _PA (tptr (Tstruct _bandmat_t noattr)))
          (Tstruct _bandmat_t noattr)) _b tint))
    (Ssequence
      (Sset _k (Econst_int (Int.repr 0) tint))
      (Sloop
        (Ssequence
          (Sifthenelse (Ebinop Olt (Etempvar _k tint) (Etempvar _n tint)
                         tint)
            Sskip
            Sbreak)
          (Ssequence
            (Ssequence
              (Scall (Some _t'1)
                (Evar _bandmat_get (Tfunction
                                     ((tptr (Tstruct _bandmat_t noattr)) ::
                                      tint :: tint :: nil) tdouble
                                     cc_default))
                ((Etempvar _PA (tptr (Tstruct _bandmat_t noattr))) ::
                 (Etempvar _k tint) :: (Econst_int (Int.repr 0) tint) :: nil))
              (Sifthenelse (Ebinop Oge (Etempvar _t'1 tdouble)
                             (Econst_int (Int.repr 0) tint) tint)
                Sskip
                (Ssequence
                  (Scall None
                    (Evar _printf (Tfunction ((tptr tschar) :: nil) tint
                                    {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|}))
                    ((Evar ___stringlit_5 (tarray tschar 30)) ::
                     (Evar ___stringlit_4 (tarray tschar 10)) ::
                     (Econst_int (Int.repr 131) tint) ::
                     (Evar ___stringlit_3 (tarray tschar 25)) :: nil))
                  (Scall None (Evar _abort (Tfunction nil tvoid cc_default))
                    nil))))
            (Ssequence
              (Ssequence
                (Ssequence
                  (Scall (Some _t'2)
                    (Evar _bandmat_get (Tfunction
                                         ((tptr (Tstruct _bandmat_t noattr)) ::
                                          tint :: tint :: nil) tdouble
                                         cc_default))
                    ((Etempvar _PA (tptr (Tstruct _bandmat_t noattr))) ::
                     (Etempvar _k tint) :: (Econst_int (Int.repr 0) tint) ::
                     nil))
                  (Scall (Some _t'3)
                    (Evar _sqrt (Tfunction (tdouble :: nil) tdouble
                                  cc_default))
                    ((Etempvar _t'2 tdouble) :: nil)))
                (Scall None
                  (Evar _bandmat_set (Tfunction
                                       ((tptr (Tstruct _bandmat_t noattr)) ::
                                        tint :: tint :: tdouble :: nil) tvoid
                                       cc_default))
                  ((Etempvar _PA (tptr (Tstruct _bandmat_t noattr))) ::
                   (Etempvar _k tint) :: (Econst_int (Int.repr 0) tint) ::
                   (Etempvar _t'3 tdouble) :: nil)))
              (Ssequence
                (Ssequence
                  (Sset _j
                    (Ebinop Oadd (Etempvar _k tint)
                      (Econst_int (Int.repr 1) tint) tint))
                  (Sloop
                    (Ssequence
                      (Ssequence
                        (Sifthenelse (Ebinop Olt (Etempvar _j tint)
                                       (Etempvar _n tint) tint)
                          (Sset _t'4
                            (Ecast
                              (Ebinop Ole (Etempvar _j tint)
                                (Ebinop Oadd (Etempvar _k tint)
                                  (Etempvar _bw tint) tint) tint) tbool))
                          (Sset _t'4 (Econst_int (Int.repr 0) tint)))
                        (Sifthenelse (Etempvar _t'4 tint) Sskip Sbreak))
                      (Ssequence
                        (Ssequence
                          (Scall (Some _t'5)
                            (Evar _bandmat_get (Tfunction
                                                 ((tptr (Tstruct _bandmat_t noattr)) ::
                                                  tint :: tint :: nil)
                                                 tdouble cc_default))
                            ((Etempvar _PA (tptr (Tstruct _bandmat_t noattr))) ::
                             (Etempvar _j tint) ::
                             (Ebinop Osub (Etempvar _j tint)
                               (Etempvar _k tint) tint) :: nil))
                          (Scall (Some _t'6)
                            (Evar _bandmat_get (Tfunction
                                                 ((tptr (Tstruct _bandmat_t noattr)) ::
                                                  tint :: tint :: nil)
                                                 tdouble cc_default))
                            ((Etempvar _PA (tptr (Tstruct _bandmat_t noattr))) ::
                             (Etempvar _k tint) ::
                             (Econst_int (Int.repr 0) tint) :: nil)))
                        (Scall None
                          (Evar _bandmat_set (Tfunction
                                               ((tptr (Tstruct _bandmat_t noattr)) ::
                                                tint :: tint :: tdouble ::
                                                nil) tvoid cc_default))
                          ((Etempvar _PA (tptr (Tstruct _bandmat_t noattr))) ::
                           (Etempvar _j tint) ::
                           (Ebinop Osub (Etempvar _j tint) (Etempvar _k tint)
                             tint) ::
                           (Ebinop Odiv (Etempvar _t'5 tdouble)
                             (Etempvar _t'6 tdouble) tdouble) :: nil))))
                    (Sset _j
                      (Ebinop Oadd (Etempvar _j tint)
                        (Econst_int (Int.repr 1) tint) tint))))
                (Ssequence
                  (Sset _j__1
                    (Ebinop Oadd (Etempvar _k tint)
                      (Econst_int (Int.repr 1) tint) tint))
                  (Sloop
                    (Ssequence
                      (Ssequence
                        (Sifthenelse (Ebinop Olt (Etempvar _j__1 tint)
                                       (Etempvar _n tint) tint)
                          (Sset _t'7
                            (Ecast
                              (Ebinop Ole (Etempvar _j__1 tint)
                                (Ebinop Oadd (Etempvar _k tint)
                                  (Etempvar _bw tint) tint) tint) tbool))
                          (Sset _t'7 (Econst_int (Int.repr 0) tint)))
                        (Sifthenelse (Etempvar _t'7 tint) Sskip Sbreak))
                      (Ssequence
                        (Sset _i
                          (Ebinop Oadd (Etempvar _k tint)
                            (Econst_int (Int.repr 1) tint) tint))
                        (Sloop
                          (Ssequence
                            (Sifthenelse (Ebinop Ole (Etempvar _i tint)
                                           (Etempvar _j__1 tint) tint)
                              Sskip
                              Sbreak)
                            (Ssequence
                              (Ssequence
                                (Scall (Some _t'8)
                                  (Evar _bandmat_get (Tfunction
                                                       ((tptr (Tstruct _bandmat_t noattr)) ::
                                                        tint :: tint :: nil)
                                                       tdouble cc_default))
                                  ((Etempvar _PA (tptr (Tstruct _bandmat_t noattr))) ::
                                   (Etempvar _i tint) ::
                                   (Ebinop Osub (Etempvar _i tint)
                                     (Etempvar _k tint) tint) :: nil))
                                (Scall (Some _t'9)
                                  (Evar _bandmat_get (Tfunction
                                                       ((tptr (Tstruct _bandmat_t noattr)) ::
                                                        tint :: tint :: nil)
                                                       tdouble cc_default))
                                  ((Etempvar _PA (tptr (Tstruct _bandmat_t noattr))) ::
                                   (Etempvar _j__1 tint) ::
                                   (Ebinop Osub (Etempvar _j__1 tint)
                                     (Etempvar _k tint) tint) :: nil)))
                              (Scall None
                                (Evar _bandmat_addto (Tfunction
                                                       ((tptr (Tstruct _bandmat_t noattr)) ::
                                                        tint :: tint ::
                                                        tdouble :: nil) tvoid
                                                       cc_default))
                                ((Etempvar _PA (tptr (Tstruct _bandmat_t noattr))) ::
                                 (Etempvar _j__1 tint) ::
                                 (Ebinop Osub (Etempvar _j__1 tint)
                                   (Etempvar _i tint) tint) ::
                                 (Ebinop Omul
                                   (Eunop Oneg (Etempvar _t'8 tdouble)
                                     tdouble) (Etempvar _t'9 tdouble)
                                   tdouble) :: nil))))
                          (Sset _i
                            (Ebinop Oadd (Etempvar _i tint)
                              (Econst_int (Int.repr 1) tint) tint)))))
                    (Sset _j__1
                      (Ebinop Oadd (Etempvar _j__1 tint)
                        (Econst_int (Int.repr 1) tint) tint))))))))
        (Sset _k
          (Ebinop Oadd (Etempvar _k tint) (Econst_int (Int.repr 1) tint)
            tint))))))
|}.

Definition f_bandmat_solve := {|
  fn_return := tvoid;
  fn_callconv := cc_default;
  fn_params := ((_PR, (tptr (Tstruct _bandmat_t noattr))) ::
                (_x, (tptr tdouble)) :: nil);
  fn_vars := nil;
  fn_temps := ((_n, tint) :: (_bw, tint) :: (_i, tint) :: (_bi, tdouble) ::
               (_dj, tint) :: (_i__1, tint) :: (_yi, tdouble) ::
               (_j, tint) :: (_t'6, tdouble) :: (_t'5, tdouble) ::
               (_t'4, tint) :: (_t'3, tdouble) :: (_t'2, tdouble) ::
               (_t'1, tint) :: (_t'8, tdouble) :: (_t'7, tdouble) :: nil);
  fn_body :=
(Ssequence
  (Sset _n
    (Efield
      (Ederef (Etempvar _PR (tptr (Tstruct _bandmat_t noattr)))
        (Tstruct _bandmat_t noattr)) _m tint))
  (Ssequence
    (Sset _bw
      (Efield
        (Ederef (Etempvar _PR (tptr (Tstruct _bandmat_t noattr)))
          (Tstruct _bandmat_t noattr)) _b tint))
    (Ssequence
      (Ssequence
        (Sset _i (Econst_int (Int.repr 0) tint))
        (Sloop
          (Ssequence
            (Sifthenelse (Ebinop Olt (Etempvar _i tint) (Etempvar _n tint)
                           tint)
              Sskip
              Sbreak)
            (Ssequence
              (Sset _bi
                (Ederef
                  (Ebinop Oadd (Etempvar _x (tptr tdouble))
                    (Etempvar _i tint) (tptr tdouble)) tdouble))
              (Ssequence
                (Ssequence
                  (Sset _dj (Econst_int (Int.repr 1) tint))
                  (Sloop
                    (Ssequence
                      (Ssequence
                        (Sifthenelse (Ebinop Ole (Etempvar _dj tint)
                                       (Etempvar _bw tint) tint)
                          (Sset _t'1
                            (Ecast
                              (Ebinop Ole (Etempvar _dj tint)
                                (Etempvar _i tint) tint) tbool))
                          (Sset _t'1 (Econst_int (Int.repr 0) tint)))
                        (Sifthenelse (Etempvar _t'1 tint) Sskip Sbreak))
                      (Ssequence
                        (Scall (Some _t'2)
                          (Evar _bandmat_get (Tfunction
                                               ((tptr (Tstruct _bandmat_t noattr)) ::
                                                tint :: tint :: nil) tdouble
                                               cc_default))
                          ((Etempvar _PR (tptr (Tstruct _bandmat_t noattr))) ::
                           (Etempvar _i tint) :: (Etempvar _dj tint) :: nil))
                        (Ssequence
                          (Sset _t'8
                            (Ederef
                              (Ebinop Oadd (Etempvar _x (tptr tdouble))
                                (Ebinop Osub (Etempvar _i tint)
                                  (Etempvar _dj tint) tint) (tptr tdouble))
                              tdouble))
                          (Sset _bi
                            (Ebinop Osub (Etempvar _bi tdouble)
                              (Ebinop Omul (Etempvar _t'2 tdouble)
                                (Etempvar _t'8 tdouble) tdouble) tdouble)))))
                    (Sset _dj
                      (Ebinop Oadd (Etempvar _dj tint)
                        (Econst_int (Int.repr 1) tint) tint))))
                (Ssequence
                  (Scall (Some _t'3)
                    (Evar _bandmat_get (Tfunction
                                         ((tptr (Tstruct _bandmat_t noattr)) ::
                                          tint :: tint :: nil) tdouble
                                         cc_default))
                    ((Etempvar _PR (tptr (Tstruct _bandmat_t noattr))) ::
                     (Etempvar _i tint) :: (Econst_int (Int.repr 0) tint) ::
                     nil))
                  (Sassign
                    (Ederef
                      (Ebinop Oadd (Etempvar _x (tptr tdouble))
                        (Etempvar _i tint) (tptr tdouble)) tdouble)
                    (Ebinop Odiv (Etempvar _bi tdouble)
                      (Etempvar _t'3 tdouble) tdouble))))))
          (Sset _i
            (Ebinop Oadd (Etempvar _i tint) (Econst_int (Int.repr 1) tint)
              tint))))
      (Ssequence
        (Sset _i__1
          (Ebinop Osub (Etempvar _n tint) (Econst_int (Int.repr 1) tint)
            tint))
        (Sloop
          (Ssequence
            (Sifthenelse (Ebinop Oge (Etempvar _i__1 tint)
                           (Econst_int (Int.repr 0) tint) tint)
              Sskip
              Sbreak)
            (Ssequence
              (Sset _yi
                (Ederef
                  (Ebinop Oadd (Etempvar _x (tptr tdouble))
                    (Etempvar _i__1 tint) (tptr tdouble)) tdouble))
              (Ssequence
                (Ssequence
                  (Sset _j
                    (Ebinop Oadd (Etempvar _i__1 tint)
                      (Econst_int (Int.repr 1) tint) tint))
                  (Sloop
                    (Ssequence
                      (Ssequence
                        (Sifthenelse (Ebinop Ole (Etempvar _j tint)
                                       (Ebinop Oadd (Etempvar _i__1 tint)
                                         (Etempvar _bw tint) tint) tint)
                          (Sset _t'4
                            (Ecast
                              (Ebinop Olt (Etempvar _j tint)
                                (Etempvar _n tint) tint) tbool))
                          (Sset _t'4 (Econst_int (Int.repr 0) tint)))
                        (Sifthenelse (Etempvar _t'4 tint) Sskip Sbreak))
                      (Ssequence
                        (Scall (Some _t'5)
                          (Evar _bandmat_get (Tfunction
                                               ((tptr (Tstruct _bandmat_t noattr)) ::
                                                tint :: tint :: nil) tdouble
                                               cc_default))
                          ((Etempvar _PR (tptr (Tstruct _bandmat_t noattr))) ::
                           (Etempvar _j tint) ::
                           (Ebinop Osub (Etempvar _j tint)
                             (Etempvar _i__1 tint) tint) :: nil))
                        (Ssequence
                          (Sset _t'7
                            (Ederef
                              (Ebinop Oadd (Etempvar _x (tptr tdouble))
                                (Etempvar _j tint) (tptr tdouble)) tdouble))
                          (Sset _yi
                            (Ebinop Osub (Etempvar _yi tdouble)
                              (Ebinop Omul (Etempvar _t'5 tdouble)
                                (Etempvar _t'7 tdouble) tdouble) tdouble)))))
                    (Sset _j
                      (Ebinop Oadd (Etempvar _j tint)
                        (Econst_int (Int.repr 1) tint) tint))))
                (Ssequence
                  (Scall (Some _t'6)
                    (Evar _bandmat_get (Tfunction
                                         ((tptr (Tstruct _bandmat_t noattr)) ::
                                          tint :: tint :: nil) tdouble
                                         cc_default))
                    ((Etempvar _PR (tptr (Tstruct _bandmat_t noattr))) ::
                     (Etempvar _i__1 tint) ::
                     (Econst_int (Int.repr 0) tint) :: nil))
                  (Sassign
                    (Ederef
                      (Ebinop Oadd (Etempvar _x (tptr tdouble))
                        (Etempvar _i__1 tint) (tptr tdouble)) tdouble)
                    (Ebinop Odiv (Etempvar _yi tdouble)
                      (Etempvar _t'6 tdouble) tdouble))))))
          (Sset _i__1
            (Ebinop Osub (Etempvar _i__1 tint) (Econst_int (Int.repr 1) tint)
              tint)))))))
|}.

Definition f_bandmat_norm2 := {|
  fn_return := tdouble;
  fn_callconv := cc_default;
  fn_params := ((_vm, (tptr (Tstruct _bandmat_t noattr))) :: nil);
  fn_vars := nil;
  fn_temps := ((_t'1, tdouble) :: (_t'3, tint) :: (_t'2, tint) :: nil);
  fn_body :=
(Ssequence
  (Ssequence
    (Sset _t'2
      (Efield
        (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
          (Tstruct _bandmat_t noattr)) _m tint))
    (Ssequence
      (Sset _t'3
        (Efield
          (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
            (Tstruct _bandmat_t noattr)) _b tint))
      (Scall (Some _t'1)
        (Evar _data_norm2 (Tfunction ((tptr tdouble) :: tint :: nil) tdouble
                            cc_default))
        ((Efield
           (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
             (Tstruct _bandmat_t noattr)) _data (tarray tdouble 0)) ::
         (Ebinop Omul (Etempvar _t'2 tint)
           (Ebinop Oadd (Etempvar _t'3 tint) (Econst_int (Int.repr 1) tint)
             tint) tint) :: nil))))
  (Sreturn (Some (Etempvar _t'1 tdouble))))
|}.

Definition f_bandmat_norm := {|
  fn_return := tdouble;
  fn_callconv := cc_default;
  fn_params := ((_vm, (tptr (Tstruct _bandmat_t noattr))) :: nil);
  fn_vars := nil;
  fn_temps := ((_t'1, tdouble) :: (_t'3, tint) :: (_t'2, tint) :: nil);
  fn_body :=
(Ssequence
  (Ssequence
    (Sset _t'2
      (Efield
        (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
          (Tstruct _bandmat_t noattr)) _m tint))
    (Ssequence
      (Sset _t'3
        (Efield
          (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
            (Tstruct _bandmat_t noattr)) _b tint))
      (Scall (Some _t'1)
        (Evar _data_norm (Tfunction ((tptr tdouble) :: tint :: nil) tdouble
                           cc_default))
        ((Efield
           (Ederef (Etempvar _vm (tptr (Tstruct _bandmat_t noattr)))
             (Tstruct _bandmat_t noattr)) _data (tarray tdouble 0)) ::
         (Ebinop Omul (Etempvar _t'2 tint)
           (Ebinop Oadd (Etempvar _t'3 tint) (Econst_int (Int.repr 1) tint)
             tint) tint) :: nil))))
  (Sreturn (Some (Etempvar _t'1 tdouble))))
|}.

Definition composites : list composite_definition :=
(Composite _densemat_t Struct
   (Member_plain _m tint :: Member_plain _n tint ::
    Member_plain _data (tarray tdouble 0) :: nil)
   noattr ::
 Composite _bandmat_t Struct
   (Member_plain _m tint :: Member_plain _b tint ::
    Member_plain _data (tarray tdouble 0) :: nil)
   noattr :: nil).

Definition global_definitions : list (ident × globdef fundef type) :=
((___compcert_va_int32,
   Gfun(External (EF_runtime "__compcert_va_int32"
                   (mksignature (AST.Xptr :: nil) AST.Xint cc_default))
     ((tptr tvoid) :: nil) tuint cc_default)) ::
 (___compcert_va_int64,
   Gfun(External (EF_runtime "__compcert_va_int64"
                   (mksignature (AST.Xptr :: nil) AST.Xlong cc_default))
     ((tptr tvoid) :: nil) tulong cc_default)) ::
 (___compcert_va_float64,
   Gfun(External (EF_runtime "__compcert_va_float64"
                   (mksignature (AST.Xptr :: nil) AST.Xfloat cc_default))
     ((tptr tvoid) :: nil) tdouble cc_default)) ::
 (___compcert_va_composite,
   Gfun(External (EF_runtime "__compcert_va_composite"
                   (mksignature (AST.Xptr :: AST.Xlong :: nil) AST.Xptr
                     cc_default)) ((tptr tvoid) :: tulong :: nil)
     (tptr tvoid) cc_default)) ::
 (___compcert_i64_dtos,
   Gfun(External (EF_runtime "__compcert_i64_dtos"
                   (mksignature (AST.Xfloat :: nil) AST.Xlong cc_default))
     (tdouble :: nil) tlong cc_default)) ::
 (___compcert_i64_dtou,
   Gfun(External (EF_runtime "__compcert_i64_dtou"
                   (mksignature (AST.Xfloat :: nil) AST.Xlong cc_default))
     (tdouble :: nil) tulong cc_default)) ::
 (___compcert_i64_stod,
   Gfun(External (EF_runtime "__compcert_i64_stod"
                   (mksignature (AST.Xlong :: nil) AST.Xfloat cc_default))
     (tlong :: nil) tdouble cc_default)) ::
 (___compcert_i64_utod,
   Gfun(External (EF_runtime "__compcert_i64_utod"
                   (mksignature (AST.Xlong :: nil) AST.Xfloat cc_default))
     (tulong :: nil) tdouble cc_default)) ::
 (___compcert_i64_stof,
   Gfun(External (EF_runtime "__compcert_i64_stof"
                   (mksignature (AST.Xlong :: nil) AST.Xsingle cc_default))
     (tlong :: nil) tfloat cc_default)) ::
 (___compcert_i64_utof,
   Gfun(External (EF_runtime "__compcert_i64_utof"
                   (mksignature (AST.Xlong :: nil) AST.Xsingle cc_default))
     (tulong :: nil) tfloat cc_default)) ::
 (___compcert_i64_sdiv,
   Gfun(External (EF_runtime "__compcert_i64_sdiv"
                   (mksignature (AST.Xlong :: AST.Xlong :: nil) AST.Xlong
                     cc_default)) (tlong :: tlong :: nil) tlong cc_default)) ::
 (___compcert_i64_udiv,
   Gfun(External (EF_runtime "__compcert_i64_udiv"
                   (mksignature (AST.Xlong :: AST.Xlong :: nil) AST.Xlong
                     cc_default)) (tulong :: tulong :: nil) tulong
     cc_default)) ::
 (___compcert_i64_smod,
   Gfun(External (EF_runtime "__compcert_i64_smod"
                   (mksignature (AST.Xlong :: AST.Xlong :: nil) AST.Xlong
                     cc_default)) (tlong :: tlong :: nil) tlong cc_default)) ::
 (___compcert_i64_umod,
   Gfun(External (EF_runtime "__compcert_i64_umod"
                   (mksignature (AST.Xlong :: AST.Xlong :: nil) AST.Xlong
                     cc_default)) (tulong :: tulong :: nil) tulong
     cc_default)) ::
 (___compcert_i64_shl,
   Gfun(External (EF_runtime "__compcert_i64_shl"
                   (mksignature (AST.Xlong :: AST.Xint :: nil) AST.Xlong
                     cc_default)) (tlong :: tint :: nil) tlong cc_default)) ::
 (___compcert_i64_shr,
   Gfun(External (EF_runtime "__compcert_i64_shr"
                   (mksignature (AST.Xlong :: AST.Xint :: nil) AST.Xlong
                     cc_default)) (tulong :: tint :: nil) tulong cc_default)) ::
 (___compcert_i64_sar,
   Gfun(External (EF_runtime "__compcert_i64_sar"
                   (mksignature (AST.Xlong :: AST.Xint :: nil) AST.Xlong
                     cc_default)) (tlong :: tint :: nil) tlong cc_default)) ::
 (___compcert_i64_smulh,
   Gfun(External (EF_runtime "__compcert_i64_smulh"
                   (mksignature (AST.Xlong :: AST.Xlong :: nil) AST.Xlong
                     cc_default)) (tlong :: tlong :: nil) tlong cc_default)) ::
 (___compcert_i64_umulh,
   Gfun(External (EF_runtime "__compcert_i64_umulh"
                   (mksignature (AST.Xlong :: AST.Xlong :: nil) AST.Xlong
                     cc_default)) (tulong :: tulong :: nil) tulong
     cc_default)) :: (___stringlit_3, Gvar v___stringlit_3) ::
 (___stringlit_1, Gvar v___stringlit_1) ::
 (___stringlit_5, Gvar v___stringlit_5) ::
 (___stringlit_2, Gvar v___stringlit_2) ::
 (___stringlit_4, Gvar v___stringlit_4) ::
 (___builtin_bswap64,
   Gfun(External (EF_builtin "__builtin_bswap64"
                   (mksignature (AST.Xlong :: nil) AST.Xlong cc_default))
     (tulong :: nil) tulong cc_default)) ::
 (___builtin_bswap,
   Gfun(External (EF_builtin "__builtin_bswap"
                   (mksignature (AST.Xint :: nil) AST.Xint cc_default))
     (tuint :: nil) tuint cc_default)) ::
 (___builtin_bswap32,
   Gfun(External (EF_builtin "__builtin_bswap32"
                   (mksignature (AST.Xint :: nil) AST.Xint cc_default))
     (tuint :: nil) tuint cc_default)) ::
 (___builtin_bswap16,
   Gfun(External (EF_builtin "__builtin_bswap16"
                   (mksignature (AST.Xint16unsigned :: nil)
                     AST.Xint16unsigned cc_default)) (tushort :: nil) tushort
     cc_default)) ::
 (___builtin_clz,
   Gfun(External (EF_builtin "__builtin_clz"
                   (mksignature (AST.Xint :: nil) AST.Xint cc_default))
     (tuint :: nil) tint cc_default)) ::
 (___builtin_clzl,
   Gfun(External (EF_builtin "__builtin_clzl"
                   (mksignature (AST.Xlong :: nil) AST.Xint cc_default))
     (tulong :: nil) tint cc_default)) ::
 (___builtin_clzll,
   Gfun(External (EF_builtin "__builtin_clzll"
                   (mksignature (AST.Xlong :: nil) AST.Xint cc_default))
     (tulong :: nil) tint cc_default)) ::
 (___builtin_ctz,
   Gfun(External (EF_builtin "__builtin_ctz"
                   (mksignature (AST.Xint :: nil) AST.Xint cc_default))
     (tuint :: nil) tint cc_default)) ::
 (___builtin_ctzl,
   Gfun(External (EF_builtin "__builtin_ctzl"
                   (mksignature (AST.Xlong :: nil) AST.Xint cc_default))
     (tulong :: nil) tint cc_default)) ::
 (___builtin_ctzll,
   Gfun(External (EF_builtin "__builtin_ctzll"
                   (mksignature (AST.Xlong :: nil) AST.Xint cc_default))
     (tulong :: nil) tint cc_default)) ::
 (___builtin_fabs,
   Gfun(External (EF_builtin "__builtin_fabs"
                   (mksignature (AST.Xfloat :: nil) AST.Xfloat cc_default))
     (tdouble :: nil) tdouble cc_default)) ::
 (___builtin_fabsf,
   Gfun(External (EF_builtin "__builtin_fabsf"
                   (mksignature (AST.Xsingle :: nil) AST.Xsingle cc_default))
     (tfloat :: nil) tfloat cc_default)) ::
 (___builtin_fsqrt,
   Gfun(External (EF_builtin "__builtin_fsqrt"
                   (mksignature (AST.Xfloat :: nil) AST.Xfloat cc_default))
     (tdouble :: nil) tdouble cc_default)) ::
 (___builtin_sqrt,
   Gfun(External (EF_builtin "__builtin_sqrt"
                   (mksignature (AST.Xfloat :: nil) AST.Xfloat cc_default))
     (tdouble :: nil) tdouble cc_default)) ::
 (___builtin_memcpy_aligned,
   Gfun(External (EF_builtin "__builtin_memcpy_aligned"
                   (mksignature
                     (AST.Xptr :: AST.Xptr :: AST.Xlong :: AST.Xlong :: nil)
                     AST.Xvoid cc_default))
     ((tptr tvoid) :: (tptr tvoid) :: tulong :: tulong :: nil) tvoid
     cc_default)) ::
 (___builtin_sel,
   Gfun(External (EF_builtin "__builtin_sel"
                   (mksignature (AST.Xbool :: nil) AST.Xvoid
                     {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|}))
     (tbool :: nil) tvoid
     {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|})) ::
 (___builtin_annot,
   Gfun(External (EF_builtin "__builtin_annot"
                   (mksignature (AST.Xptr :: nil) AST.Xvoid
                     {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|}))
     ((tptr tschar) :: nil) tvoid
     {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|})) ::
 (___builtin_annot_intval,
   Gfun(External (EF_builtin "__builtin_annot_intval"
                   (mksignature (AST.Xptr :: AST.Xint :: nil) AST.Xint
                     cc_default)) ((tptr tschar) :: tint :: nil) tint
     cc_default)) ::
 (___builtin_membar,
   Gfun(External (EF_builtin "__builtin_membar"
                   (mksignature nil AST.Xvoid cc_default)) nil tvoid
     cc_default)) ::
 (___builtin_va_start,
   Gfun(External (EF_builtin "__builtin_va_start"
                   (mksignature (AST.Xptr :: nil) AST.Xvoid cc_default))
     ((tptr tvoid) :: nil) tvoid cc_default)) ::
 (___builtin_va_arg,
   Gfun(External (EF_builtin "__builtin_va_arg"
                   (mksignature (AST.Xptr :: AST.Xint :: nil) AST.Xvoid
                     cc_default)) ((tptr tvoid) :: tuint :: nil) tvoid
     cc_default)) ::
 (___builtin_va_copy,
   Gfun(External (EF_builtin "__builtin_va_copy"
                   (mksignature (AST.Xptr :: AST.Xptr :: nil) AST.Xvoid
                     cc_default)) ((tptr tvoid) :: (tptr tvoid) :: nil) tvoid
     cc_default)) ::
 (___builtin_va_end,
   Gfun(External (EF_builtin "__builtin_va_end"
                   (mksignature (AST.Xptr :: nil) AST.Xvoid cc_default))
     ((tptr tvoid) :: nil) tvoid cc_default)) ::
 (___builtin_unreachable,
   Gfun(External (EF_builtin "__builtin_unreachable"
                   (mksignature nil AST.Xvoid cc_default)) nil tvoid
     cc_default)) ::
 (___builtin_expect,
   Gfun(External (EF_builtin "__builtin_expect"
                   (mksignature (AST.Xlong :: AST.Xlong :: nil) AST.Xlong
                     cc_default)) (tlong :: tlong :: nil) tlong cc_default)) ::
 (___builtin_cls,
   Gfun(External (EF_builtin "__builtin_cls"
                   (mksignature (AST.Xint :: nil) AST.Xint cc_default))
     (tint :: nil) tint cc_default)) ::
 (___builtin_clsl,
   Gfun(External (EF_builtin "__builtin_clsl"
                   (mksignature (AST.Xlong :: nil) AST.Xint cc_default))
     (tlong :: nil) tint cc_default)) ::
 (___builtin_clsll,
   Gfun(External (EF_builtin "__builtin_clsll"
                   (mksignature (AST.Xlong :: nil) AST.Xint cc_default))
     (tlong :: nil) tint cc_default)) ::
 (___builtin_fmadd,
   Gfun(External (EF_builtin "__builtin_fmadd"
                   (mksignature
                     (AST.Xfloat :: AST.Xfloat :: AST.Xfloat :: nil)
                     AST.Xfloat cc_default))
     (tdouble :: tdouble :: tdouble :: nil) tdouble cc_default)) ::
 (___builtin_fmsub,
   Gfun(External (EF_builtin "__builtin_fmsub"
                   (mksignature
                     (AST.Xfloat :: AST.Xfloat :: AST.Xfloat :: nil)
                     AST.Xfloat cc_default))
     (tdouble :: tdouble :: tdouble :: nil) tdouble cc_default)) ::
 (___builtin_fnmadd,
   Gfun(External (EF_builtin "__builtin_fnmadd"
                   (mksignature
                     (AST.Xfloat :: AST.Xfloat :: AST.Xfloat :: nil)
                     AST.Xfloat cc_default))
     (tdouble :: tdouble :: tdouble :: nil) tdouble cc_default)) ::
 (___builtin_fnmsub,
   Gfun(External (EF_builtin "__builtin_fnmsub"
                   (mksignature
                     (AST.Xfloat :: AST.Xfloat :: AST.Xfloat :: nil)
                     AST.Xfloat cc_default))
     (tdouble :: tdouble :: tdouble :: nil) tdouble cc_default)) ::
 (___builtin_fmax,
   Gfun(External (EF_builtin "__builtin_fmax"
                   (mksignature (AST.Xfloat :: AST.Xfloat :: nil) AST.Xfloat
                     cc_default)) (tdouble :: tdouble :: nil) tdouble
     cc_default)) ::
 (___builtin_fmin,
   Gfun(External (EF_builtin "__builtin_fmin"
                   (mksignature (AST.Xfloat :: AST.Xfloat :: nil) AST.Xfloat
                     cc_default)) (tdouble :: tdouble :: nil) tdouble
     cc_default)) ::
 (___builtin_debug,
   Gfun(External (EF_external "__builtin_debug"
                   (mksignature (AST.Xint :: nil) AST.Xvoid
                     {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|}))
     (tint :: nil) tvoid
     {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|})) ::
 (_sqrt,
   Gfun(External (EF_external "sqrt"
                   (mksignature (AST.Xfloat :: nil) AST.Xfloat cc_default))
     (tdouble :: nil) tdouble cc_default)) ::
 (_printf,
   Gfun(External (EF_external "printf"
                   (mksignature (AST.Xptr :: nil) AST.Xint
                     {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|}))
     ((tptr tschar) :: nil) tint
     {|cc_vararg:=(Some 1); cc_unproto:=false; cc_structret:=false|})) ::
 (_free, Gfun(External EF_free ((tptr tvoid) :: nil) tvoid cc_default)) ::
 (_abort,
   Gfun(External (EF_external "abort" (mksignature nil AST.Xvoid cc_default))
     nil tvoid cc_default)) ::
 (_memset,
   Gfun(External (EF_external "memset"
                   (mksignature (AST.Xptr :: AST.Xint :: AST.Xlong :: nil)
                     AST.Xptr cc_default))
     ((tptr tvoid) :: tint :: tulong :: nil) (tptr tvoid) cc_default)) ::
 (_surely_malloc,
   Gfun(External (EF_external "surely_malloc"
                   (mksignature (AST.Xlong :: nil) AST.Xptr cc_default))
     (tulong :: nil) (tptr tvoid) cc_default)) ::
 (_data_norm2,
   Gfun(External (EF_external "data_norm2"
                   (mksignature (AST.Xptr :: AST.Xint :: nil) AST.Xfloat
                     cc_default)) ((tptr tdouble) :: tint :: nil) tdouble
     cc_default)) ::
 (_data_norm,
   Gfun(External (EF_external "data_norm"
                   (mksignature (AST.Xptr :: AST.Xint :: nil) AST.Xfloat
                     cc_default)) ((tptr tdouble) :: tint :: nil) tdouble
     cc_default)) ::
 (_densemat_get,
   Gfun(External (EF_external "densemat_get"
                   (mksignature (AST.Xptr :: AST.Xint :: AST.Xint :: nil)
                     AST.Xfloat cc_default))
     ((tptr (Tstruct _densemat_t noattr)) :: tint :: tint :: nil) tdouble
     cc_default)) :: (_bandmat_malloc, Gfun(Internal f_bandmat_malloc)) ::
 (_bandmat_free, Gfun(Internal f_bandmat_free)) ::
 (_bandmatn_clear, Gfun(Internal f_bandmatn_clear)) ::
 (_bandmat_clear, Gfun(Internal f_bandmat_clear)) ::
 (_bandmatn_get, Gfun(Internal f_bandmatn_get)) ::
 (_bandmatn_set, Gfun(Internal f_bandmatn_set)) ::
 (_bandmatn_addto, Gfun(Internal f_bandmatn_addto)) ::
 (_bandmat_get, Gfun(Internal f_bandmat_get)) ::
 (_bandmat_set, Gfun(Internal f_bandmat_set)) ::
 (_bandmat_addto, Gfun(Internal f_bandmat_addto)) ::
 (_dense_to_band, Gfun(Internal f_dense_to_band)) ::
 (_bandmat_print, Gfun(Internal f_bandmat_print)) ::
 (_bandmat_factor, Gfun(Internal f_bandmat_factor)) ::
 (_bandmat_solve, Gfun(Internal f_bandmat_solve)) ::
 (_bandmat_norm2, Gfun(Internal f_bandmat_norm2)) ::
 (_bandmat_norm, Gfun(Internal f_bandmat_norm)) :: nil).

Definition public_idents : list ident :=
(_bandmat_norm :: _bandmat_norm2 :: _bandmat_solve :: _bandmat_factor ::
 _bandmat_print :: _dense_to_band :: _bandmat_addto :: _bandmat_set ::
 _bandmat_get :: _bandmatn_addto :: _bandmatn_set :: _bandmatn_get ::
 _bandmat_clear :: _bandmatn_clear :: _bandmat_free :: _bandmat_malloc ::
 _densemat_get :: _data_norm :: _data_norm2 :: _surely_malloc :: _memset ::
 _abort :: _free :: _printf :: _sqrt :: ___builtin_debug ::
 ___builtin_fmin :: ___builtin_fmax :: ___builtin_fnmsub ::
 ___builtin_fnmadd :: ___builtin_fmsub :: ___builtin_fmadd ::
 ___builtin_clsll :: ___builtin_clsl :: ___builtin_cls ::
 ___builtin_expect :: ___builtin_unreachable :: ___builtin_va_end ::
 ___builtin_va_copy :: ___builtin_va_arg :: ___builtin_va_start ::
 ___builtin_membar :: ___builtin_annot_intval :: ___builtin_annot ::
 ___builtin_sel :: ___builtin_memcpy_aligned :: ___builtin_sqrt ::
 ___builtin_fsqrt :: ___builtin_fabsf :: ___builtin_fabs ::
 ___builtin_ctzll :: ___builtin_ctzl :: ___builtin_ctz :: ___builtin_clzll ::
 ___builtin_clzl :: ___builtin_clz :: ___builtin_bswap16 ::
 ___builtin_bswap32 :: ___builtin_bswap :: ___builtin_bswap64 ::
 ___compcert_i64_umulh :: ___compcert_i64_smulh :: ___compcert_i64_sar ::
 ___compcert_i64_shr :: ___compcert_i64_shl :: ___compcert_i64_umod ::
 ___compcert_i64_smod :: ___compcert_i64_udiv :: ___compcert_i64_sdiv ::
 ___compcert_i64_utof :: ___compcert_i64_stof :: ___compcert_i64_utod ::
 ___compcert_i64_stod :: ___compcert_i64_dtou :: ___compcert_i64_dtos ::
 ___compcert_va_composite :: ___compcert_va_float64 ::
 ___compcert_va_int64 :: ___compcert_va_int32 :: nil).

Definition prog : Clight.program :=
  mkprogram composites global_definitions public_idents _main Logic.I.