PTAllocCode

Require Import Coqlib.
Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Cop.
Require Import Clight.
Require Import CDataTypes.
Require Import Ctypes.

Require Import Ident.

Local Open Scope Z_scope.

Definition _Rd : ident := 999%positive.
Definition _addr := 1%positive.
Definition _alloc := 2%positive.
Definition _arg := 3%positive.
Definition _base := 4%positive.
Definition _cb_offset := 5%positive.
Definition _cbndx := 6%positive.
Definition _cnt := 7%positive.
Definition _count := 8%positive.
Definition _cur_vcpuid := 9%positive.
Definition _cur_vmid := 10%positive.
Definition _data := 11%positive.
Definition _end := 12%positive.
Definition _esr := 13%positive.
Definition _fault_ipa := 14%positive.
Definition _gfn := 15%positive.
Definition _gpa := 16%positive.
Definition _hsr := 17%positive.
Definition _i := 18%positive.
Definition _index := 19%positive.
Definition _iova := 20%positive.
Definition _is_write := 21%positive.
Definition _kvm := 22%positive.
Definition _len := 23%positive.
Definition _level := 24%positive.
Definition _load_addr := 25%positive.
Definition _load_idx := 26%positive.
Definition _load_info_cnt := 27%positive.
Definition _main := 28%positive.
Definition _map := 29%positive.
Definition _mapped := 30%positive.
Definition _n := 31%positive.
Definition _next := 32%positive.
Definition _num := 33%positive.
Definition _num_context_banks := 34%positive.
Definition _offset := 35%positive.
Definition _owner := 36%positive.
Definition _pa := 37%positive.
Definition _paddr := 38%positive.
Definition _page_count := 39%positive.
Definition _pass_lock := 40%positive.
Definition _perm := 41%positive.
Definition _pfn := 42%positive.
Definition _pgd := 43%positive.
Definition _pgd_idx := 44%positive.
Definition _pgd_pa := 45%positive.
Definition _pgnum := 46%positive.
Definition _pmd := 47%positive.
Definition _pmd_idx := 48%positive.
Definition _pmd_pa := 49%positive.
Definition _power := 50%positive.
Definition _prot := 51%positive.
Definition _pte := 52%positive.
Definition _pte_idx := 53%positive.
Definition _pte_pa := 54%positive.
Definition _pud := 55%positive.
Definition _reg := 56%positive.
Definition _remap := 57%positive.
Definition _remap_addr := 58%positive.
Definition _res := 59%positive.
Definition _ret := 60%positive.
Definition _rt := 61%positive.
Definition _size := 62%positive.
Definition _smmu_enable := 63%positive.
Definition _smmu_index := 64%positive.
Definition _start := 65%positive.
Definition _state := 66%positive.
Definition _t_vmid := 67%positive.
Definition _target := 68%positive.
Definition _target_addr := 69%positive.
Definition _target_vmid := 70%positive.
Definition _total_smmu := 71%positive.
Definition _ttbr := 72%positive.
Definition _ttbr_pa := 73%positive.
Definition _type := 74%positive.
Definition _val := 75%positive.
Definition _valid := 76%positive.
Definition _vcpu := 77%positive.
Definition _vcpu_state := 78%positive.
Definition _vcpuid := 79%positive.
Definition _vm_state := 80%positive.
Definition _vmid := 81%positive.
Definition _vttbr := 82%positive.
Definition _wait_lock := 83%positive.
Definition _write_val := 84%positive.
Definition _t'1 := 85%positive.
Definition _t'2 := 86%positive.

Definition alloc_pgd_page_body :=
  (Ssequence
    (Ssequence
      (Scall (Some _t'1)
        (Evar get_pgd_next (Tfunction (Tcons tuint Tnil) tulong cc_default))
        ((Etempvar _vmid tuint) :: nil))
      (Sset _next (Etempvar _t'1 tulong)))
    (Ssequence
      (Sset _end
        (Ebinop Oadd
          (Ebinop Oadd (Econst_long (Int64.repr 65536) tulong)
            (Ebinop Omul (Econst_long (Int64.repr 33554432) tulong)
              (Etempvar _vmid tuint) tulong) tulong)
          (Econst_long (Int64.repr 1052672) tulong) tulong))
      (Ssequence
        (Sifthenelse (Ebinop Ole
                       (Ebinop Oadd (Etempvar _next tulong)
                         (Econst_long (Int64.repr 4096) tulong) tulong)
                       (Etempvar _end tulong) tint)
          (Scall None
            (Evar set_pgd_next (Tfunction (Tcons tuint (Tcons tulong Tnil))
                                  tvoid cc_default))
            ((Etempvar _vmid tuint) ::
             (Ebinop Oadd (Etempvar _next tulong)
               (Econst_long (Int64.repr 4096) tulong) tulong) :: nil))
          (Scall None (Evar panic (Tfunction Tnil tvoid cc_default)) nil))
        (Ssequence
          (Scall (Some _t'2)
            (Evar check64 (Tfunction (Tcons tulong Tnil) tulong cc_default))
            ((Etempvar _next tulong) :: nil))
          (Sreturn (Some (Etempvar _t'2 tulong))))))).

Definition f_alloc_pgd_page := {|
  fn_return := tulong;
  fn_callconv := cc_default;
  fn_params := ((_vmid, tuint) :: nil);
  fn_vars := nil;
  fn_temps := ((_next, tulong) :: (_end, tulong) :: (_t'2, tulong) ::
               (_t'1, tulong) :: nil);
  fn_body := alloc_pgd_page_body
|}.

Definition alloc_pud_page_body :=
  (Ssequence
    (Ssequence
      (Scall (Some _t'1)
        (Evar get_pud_next (Tfunction (Tcons tuint Tnil) tulong cc_default))
        ((Etempvar _vmid tuint) :: nil))
      (Sset _next (Etempvar _t'1 tulong)))
    (Ssequence
      (Sset _end
        (Ebinop Oadd
          (Ebinop Oadd (Econst_long (Int64.repr 65536) tulong)
            (Ebinop Omul (Econst_long (Int64.repr 33554432) tulong)
              (Etempvar _vmid tuint) tulong) tulong)
          (Econst_long (Int64.repr 9441280) tulong) tulong))
      (Ssequence
        (Sifthenelse (Ebinop Ole
                       (Ebinop Oadd (Etempvar _next tulong)
                         (Econst_long (Int64.repr 4096) tulong) tulong)
                       (Etempvar _end tulong) tint)
          (Scall None
            (Evar set_pud_next (Tfunction (Tcons tuint (Tcons tulong Tnil))
                                  tvoid cc_default))
            ((Etempvar _vmid tuint) ::
             (Ebinop Oadd (Etempvar _next tulong)
               (Econst_long (Int64.repr 4096) tulong) tulong) :: nil))
          (Scall None (Evar panic (Tfunction Tnil tvoid cc_default)) nil))
        (Ssequence
          (Scall (Some _t'2)
            (Evar check64 (Tfunction (Tcons tulong Tnil) tulong cc_default))
            ((Etempvar _next tulong) :: nil))
          (Sreturn (Some (Etempvar _t'2 tulong))))))).

Definition f_alloc_pud_page := {|
  fn_return := tulong;
  fn_callconv := cc_default;
  fn_params := ((_vmid, tuint) :: nil);
  fn_vars := nil;
  fn_temps := ((_next, tulong) :: (_end, tulong) :: (_t'2, tulong) ::
               (_t'1, tulong) :: nil);
  fn_body := alloc_pud_page_body
|}.

Definition alloc_pmd_page_body :=
  (Ssequence
    (Ssequence
      (Scall (Some _t'1)
        (Evar get_pmd_next (Tfunction (Tcons tuint Tnil) tulong cc_default))
        ((Etempvar _vmid tuint) :: nil))
      (Sset _next (Etempvar _t'1 tulong)))
    (Ssequence
      (Sset _end
        (Ebinop Oadd (Econst_long (Int64.repr 65536) tulong)
          (Ebinop Omul (Econst_long (Int64.repr 33554432) tulong)
            (Ebinop Oadd (Etempvar _vmid tuint) (Econst_int (Int.repr 1) tint)
              tuint) tulong) tulong))
      (Ssequence
        (Sifthenelse (Ebinop Ole
                       (Ebinop Oadd (Etempvar _next tulong)
                         (Econst_long (Int64.repr 4096) tulong) tulong)
                       (Etempvar _end tulong) tint)
          (Scall None
            (Evar set_pmd_next (Tfunction (Tcons tuint (Tcons tulong Tnil))
                                  tvoid cc_default))
            ((Etempvar _vmid tuint) ::
             (Ebinop Oadd (Etempvar _next tulong)
               (Econst_long (Int64.repr 4096) tulong) tulong) :: nil))
          (Scall None (Evar panic (Tfunction Tnil tvoid cc_default)) nil))
        (Ssequence
          (Scall (Some _t'2)
            (Evar check64 (Tfunction (Tcons tulong Tnil) tulong cc_default))
            ((Etempvar _next tulong) :: nil))
          (Sreturn (Some (Etempvar _t'2 tulong))))))).

Definition f_alloc_pmd_page := {|
  fn_return := tulong;
  fn_callconv := cc_default;
  fn_params := ((_vmid, tuint) :: nil);
  fn_vars := nil;
  fn_temps := ((_next, tulong) :: (_end, tulong) :: (_t'2, tulong) ::
               (_t'1, tulong) :: nil);
  fn_body := alloc_pmd_page_body
|}.