Theory Grid

section ‹ Sparse Grids ›

theory Grid
imports Grid_Point
begin

subsection "Vectors"

type_synonym vector = "grid_point  real"

definition null_vector :: "vector"
where "null_vector  λ p. 0"

definition sum_vector :: "vector  vector  vector"
where "sum_vector α β  λ p. α p + β p"

subsection ‹ Inductive enumeration of all grid points ›

inductive_set
  grid :: "grid_point  nat set  grid_point set"
  for b :: "grid_point" and ds :: "nat set"
where
    Start[intro!]: "b  grid b ds"
  | Child[intro!]: " p  grid b ds ; d  ds   child p dir d  grid b ds"

lemma grid_length[simp]: "p'  grid p ds  length p' = length p"
  by (erule grid.induct, auto)

lemma grid_union_dims: " ds  ds' ; p  grid b ds   p  grid b ds'"
  by (erule grid.induct, auto)

lemma grid_transitive: " a  grid b ds ; b  grid c ds' ; ds'  ds'' ; ds  ds''   a  grid c ds''"
  by (erule grid.induct, auto simp add: grid_union_dims)

lemma grid_child[intro?]: assumes "d  ds" and p_grid: "p  grid (child b dir d) ds"
  shows "p  grid b ds"
  using d  ds grid_transitive[OF p_grid] by auto

lemma grid_single_level[simp]: assumes "p  grid b ds" and "d < length b"
  shows "lv b d  lv p d"
  using assms
proof induct
  case (Child p' d' dir)
  thus ?case by (cases "d' = d", auto simp add: child_def ix_def lv_def)
qed auto

lemma grid_child_level:
  assumes "d < length b"
  and p_grid: "p  grid (child b dir d) ds"
  shows "lv b d < lv p d"
proof -
  have "lv b d < lv (child b dir d) d" using child_lv[OF d < length b] by auto
  also have "  lv p d" using p_grid assms by (intro grid_single_level) auto
  finally show ?thesis .
qed

lemma child_out: "length p  d  child p dir d = p"
  unfolding child_def by auto

lemma grid_dim_remove:
  assumes inset: "p  grid b ({d}  ds)"
  and eq: "d < length b  p ! d = b ! d"
  shows "p  grid b ds"
  using inset eq
proof induct
  case (Child p' d' dir)
  show ?case
  proof (cases "d'  length p'")
    case True with child_out[OF this]
    show ?thesis using Child by auto
  next
    case False hence "d' < length p'" by simp
    show ?thesis
    proof (cases "d' = d")
      case True
      hence "lv b d  lv p' d" and "lv p' d < lv (child p' dir d) d"
        using child_single_level Child d' < length p' by auto
      hence False using Child and d' = d and lv_def and ¬ d'  length p' by auto
      thus ?thesis ..
    next
      case False
      hence "d'  ds" using Child by auto
      moreover have "d < length b  p' ! d = b ! d"
      proof -
        assume "d < length b"
        hence "d < length p'" using Child by auto
        hence "child p' dir d' ! d = p' ! d" using child_invariant and False by auto
        thus ?thesis using Child and d < length b by auto
      qed
      hence "p'  grid b ds" using Child by auto
      ultimately show ?thesis using grid.Child by auto
    qed
  qed
qed auto

lemma gridgen_dim_restrict:
  assumes inset: "p  grid b (ds'  ds)"
  and eq: " d  ds'. d  length b"
  shows "p  grid b ds"
  using inset eq
proof induct
  case (Child p' d dir)
  thus ?case
  proof (cases "d  ds")
    case False thus ?thesis using Child and child_def by auto
  qed auto
qed auto

lemma grid_dim_remove_outer: "grid b ds = grid b {d  ds. d < length b}"
proof
  have "{d  ds. d < length b}  ds" by auto
  from grid_union_dims[OF this]
  show "grid b {d  ds. d < length b}  grid b ds" by auto

  have "ds = (ds - {d  ds. d < length b})  {d  ds. d < length b}" by auto
  moreover
  have "grid b ((ds - {d  ds. d < length b})  {d  ds. d < length b})  grid b {d  ds. d < length b}"
  proof
    fix p
    assume "p  grid b (ds - {d  ds. d < length b}  {d  ds. d < length b})"
    moreover have " d  (ds - {d  ds. d < length b}). d  length b" by auto
    ultimately show "p  grid b {d  ds. d < length b}" by (rule gridgen_dim_restrict)
  qed
  ultimately show "grid b ds  grid b {d  ds. d < length b}" by auto
qed

lemma grid_level[intro]: assumes "p  grid b ds" shows "level b  level p"
proof -
  have *: "length p = length b" using grid_length assms by auto
  { fix i assume "i  {0 ..< length p}"
    hence "lv b i  lv p i" using p  grid b ds and grid_single_level * by auto
  } thus ?thesis unfolding level_def * by (auto intro!: sum_mono)
qed
lemma grid_empty_ds[simp]: "grid b {} = { b }"
proof -
  have "!! z. z  grid b {}  z = b"
    by (erule grid.induct, auto)
  thus ?thesis by auto
qed
lemma grid_Start: assumes inset: "p  grid b ds" and eq: "level p = level b" shows "p = b"
  using inset eq
proof induct
  case (Child p d dir)
  show ?case
  proof (cases "d < length b")
    case True
    from Child
    have "level p  level b" by auto
    moreover
    have "level p  level (child p dir d)" by (rule child_level_gt)
    hence "level p  level b" using Child by auto
    ultimately have "level p = level b" by auto
    hence "p = b " using Child(2) by auto
    with Child(4) have "level (child b dir d) = level b" by auto
    moreover have "level (child b dir d)   level b" using child_level and d < length b by auto
    ultimately show ?thesis by auto
  next
    case False
    with Child have "length p = length b" by auto
    with False have "child p dir d = p" using child_def by auto
    moreover with Child have "level p = level b" by auto
    with Child(2) have "p = b" by auto
    ultimately show ?thesis by auto
  qed
qed auto
lemma grid_estimate:
  assumes "d < length b" and p_grid: "p  grid b ds"
  shows "ix p d < (ix b d + 1) * 2^(lv p d - lv b d)  ix p d > (ix b d - 1) * 2^(lv p d - lv b d)"
  using p_grid
proof induct
  case (Child p d' dir)
  show ?case
  proof (cases "d = d'")
    case False with Child show ?thesis unfolding child_def lv_def ix_def by auto
  next
    case True with child_estimate_child and Child and d < length b
    show ?thesis using grid_single_level by auto
  qed
qed auto
lemma grid_odd: assumes "d < length b" and p_diff: "p ! d  b ! d" and p_grid: "p  grid b ds"
  shows "odd (ix p d)"
  using p_grid and p_diff
proof induct
  case (Child p d' dir)
  show ?case
  proof (cases "d = d'")
    case True with child_odd and d < length b and Child show ?thesis by auto
  next
    case False with Child and d < length b show ?thesis using child_def and ix_def and lv_def by auto
  qed
qed auto
lemma grid_invariant: assumes "d < length b" and "d  ds" and p_grid: "p  grid b ds"
  shows "p ! d = b ! d"
  using p_grid
proof (induct)
  case (Child p d' dir) hence "d'  d" using d  ds by auto
  thus ?case using child_def and Child by auto
qed auto
lemma grid_part: assumes "d < length b" and p_valid: "p  grid b {d}" and p'_valid: "p'  grid b {d}"
  and level: "lv p' d  lv p d"
  and right: "ix p' d  (ix p d + 1) * 2^(lv p' d - lv p d)" (is "?right p p' d")
  and left: "ix p' d  (ix p d - 1) * 2^(lv p' d - lv p d)" (is "?left p p' d")
  shows "p'  grid p {d}"
  using p'_valid left right level and p_valid
proof induct
  case (Child p' d' dir)
  hence "d = d'" by auto
  let ?child = "child p' dir d'"

  show ?case
  proof (cases "lv p d = lv ?child d")
    case False
    moreover have "lv ?child d = lv p' d + 1" using child_lv and d < length b and Child and d = d' by auto
    ultimately have "lv p d < lv p' d + 1" using Child by auto
    hence lv: "Suc (lv p' d) - lv p d = Suc (lv p' d - lv p d)" by auto

    have "?left p p' d  ?right p p' d"
    proof (cases dir)
      case left
      with Child have "2 * ix p' d - 1  (ix p d + 1) * 2^(Suc (lv p' d) - lv p d)"
        using d = d' and d < length b by (auto simp add: child_def ix_def lv_def)
      also have " = 2 * (ix p d + 1) * 2^(lv p' d - lv p d)" using lv by auto
      finally have "2 * ix p' d - 2 < 2 * (ix p d + 1) * 2^(lv p' d - lv p d)" by auto
      also have " = 2 * ((ix p d + 1) * 2^(lv p' d - lv p d))" by auto
      finally have left_r: "ix p' d  (ix p d + 1) * 2^(lv p' d - lv p d)" by auto

      have "2 * ((ix p d - 1) * 2^(lv p' d - lv p d)) = 2 * (ix p d - 1) * 2^(lv p' d - lv p d)" by auto
      also have " = (ix p d - 1) * 2^(Suc (lv p' d) - lv p d)" using lv by auto
      also have "  2 * ix p' d - 1"
        using left and Child and d = d' and d < length b by (auto simp add: child_def ix_def lv_def)
      finally have right_r: "((ix p d - 1) * 2^(lv p' d - lv p d))  ix p' d" by auto

      show ?thesis using left_r and right_r by auto
    next
      case right
      with Child have "2 * ix p' d + 1  (ix p d + 1) * 2^(Suc (lv p' d) - lv p d)"
        using d = d' and d < length b by (auto simp add: child_def ix_def lv_def)
      also have " = 2 * (ix p d + 1) * 2^(lv p' d - lv p d)" using lv by auto
      finally have "2 * ix p' d < 2 * (ix p d + 1) * 2^(lv p' d - lv p d)" by auto
      also have " = 2 * ((ix p d + 1) * 2^(lv p' d - lv p d))" by auto
      finally have left_r: "ix p' d  (ix p d + 1) * 2^(lv p' d - lv p d)" by auto

      have "2 * ((ix p d - 1) * 2^(lv p' d - lv p d)) = 2 * (ix p d - 1) * 2^(lv p' d - lv p d)" by auto
      also have " = (ix p d - 1) * 2^(Suc (lv p' d) - lv p d)" using lv by auto
      also have "  2 * ix p' d + 1"
        using right and Child and d = d' and d < length b by (auto simp add: child_def ix_def lv_def)
      also have " < 2 * (ix p' d + 1)" by auto
      finally have right_r: "((ix p d - 1) * 2^(lv p' d - lv p d))  ix p' d" by auto

      show ?thesis using left_r and right_r by auto
    qed
    with Child and lv have "p'  grid p {d}" by auto
    thus ?thesis using d = d' by auto
  next
    case True
    moreover with Child have "?left p ?child d  ?right p ?child d" by auto
    ultimately have range: "ix p d - 1  ix ?child d  ix ?child d  ix p d + 1" by auto

    have "p ! d  b ! d"
    proof (rule ccontr)
      assume "¬ (p ! d  b ! d)"
      with lv p d = lv ?child d have "lv b d = lv ?child d" by (auto simp add: lv_def)
      hence "lv b d = lv p' d + 1" using d = d' and Child and d < length b and child_lv by auto
      moreover have "lv b d  lv p' d" using d = d' and Child and d < length b and grid_single_level by auto
      ultimately show False by auto
    qed
    hence "odd (ix p d)" using grid_odd and p  grid b {d} and d < length b by auto
    hence "¬ odd (ix p d + 1)" and "¬ odd (ix p d - 1)" by auto

    have "d < length p'" using p'  grid b {d} and d < length b by auto
    hence odd_child: "odd (ix ?child d)" using child_odd and d = d' by auto

    have "ix p d - 1  ix ?child d"
    proof (rule ccontr)
      assume "¬ (ix p d - 1  ix ?child d)"
      hence "odd (ix p d - 1)" using odd_child by auto
      thus False using ¬ odd (ix p d - 1) by auto
    qed
    moreover
    have "ix p d + 1  ix ?child d"
    proof (rule ccontr)
      assume "¬ (ix p d + 1  ix ?child d)"
      hence "odd (ix p d + 1)" using odd_child by auto
      thus False using ¬ odd (ix p d + 1) by auto
    qed
    ultimately have "ix p d = ix ?child d" using range by auto
    with True have d_eq: "p ! d = (?child) ! d" by (auto simp add: prod_eqI ix_def lv_def)

    have "length p = length ?child" using p  grid b {d} and p'  grid b {d} by auto
    moreover have "p ! d'' = ?child ! d''" if "d'' < length p" for d''
    proof -
      have "d'' < length b" using that p  grid b {d} by auto
      show "p ! d'' = ?child ! d''"
      proof (cases "d = d''")
        case True with d_eq show ?thesis by auto
      next
        case False hence "d''  {d}" by auto
        from d'' < length b and this and p  grid b {d}
        have "p ! d'' = b ! d''" by (rule grid_invariant)
        also have " = p' ! d''" using d'' < length b and d''  {d} and p'  grid b {d}
          by (rule grid_invariant[symmetric])
        also have " = ?child ! d''"
        proof -
          have "d'' < length p'" using d'' < length b and p'  grid b {d} by auto
          hence "?child ! d'' = p' ! d''" using child_invariant and d  d'' and d = d' by auto
          thus ?thesis by auto
        qed
        finally show ?thesis .
      qed
    qed
    ultimately have "p = ?child" by (rule nth_equalityI)
    thus "?child  grid p {d}" by auto
  qed
next
  case Start
  moreover hence "lv b d  lv p d" using grid_single_level and d < length b by auto
  ultimately have "lv b d = lv p d" by auto

  have "level p = level b"
  proof -
    { fix d'
      assume "d' < length b"
      have "lv b d' = lv p d'"
      proof (cases "d = d'")
        case True with lv b d = lv p d show ?thesis by auto
      next
        case False hence "d'  {d}" by auto
        from d' < length b and this and p  grid b {d}
        have "p ! d' = b ! d'" by (rule grid_invariant)
        thus ?thesis by (auto simp add: lv_def)
      qed }
    moreover have "length b = length p" using p  grid b {d} by auto
    ultimately show ?thesis by (rule level_all_eq)
  qed
  hence "p = b" using grid_Start and p  grid b {d} by auto
  thus ?case by auto
qed
lemma grid_disjunct: assumes "d < length p"
  shows "grid (child p left d) ds  grid (child p right d) ds = {}"
  (is "grid ?l ds  grid ?r ds = {}")
proof (intro set_eqI iffI)
  fix x
  assume "x  grid ?l ds  grid ?r ds"
  hence "ix x d < (ix ?l d + 1) * 2^(lv x d - lv ?l d)"
    and "ix x d > (ix ?r d - 1) * 2^(lv x d - lv ?r d)"
    using grid_estimate d < length p by auto
  thus "x  {}" using d < length p and child_lv and child_ix by auto
qed auto

lemma grid_level_eq: assumes eq: " d  ds. lv p d = lv b d" and grid: "p  grid b ds"
  shows "level p = level b"
proof (rule level_all_eq)
  { fix i assume "i < length b"
    show "lv b i = lv p i"
    proof (cases "i  ds")
      case True with eq show ?thesis by auto
    next case False with i < length b and grid show ?thesis
        using lv_def ix_def grid_invariant by auto
    qed }
  show "length b = length p" using grid by auto
qed

lemma grid_partition:
  "grid p {d} = {p}  grid (child p left d) {d}  grid (child p right d) {d}"
  (is "_ = _  grid ?l {d}  grid ?r {d}")
proof -
  have "!! x.  x  grid p {d} ; x  p ; x  grid ?r {d}   x  grid ?l {d}"
  proof (cases "d < length p")
    case True
    fix x

    let "?nr_r p" = "ix x d > (ix p d + 1) * 2 ^ (lv x d - lv p d)"
    let "?nr_l p" = "(ix p d - 1) * 2 ^ (lv x d - lv p d) > ix x d"

    have ix_r_eq: "ix ?r d = 2 * ix p d + 1" using d < length p and child_ix by auto
    have lv_r_eq: "lv ?r d = lv p d + 1" using d < length p and child_lv by auto

    have ix_l_eq: "ix ?l d = 2 * ix p d - 1" using d < length p and child_ix by auto
    have lv_l_eq: "lv ?l d = lv p d + 1" using d < length p and child_lv by auto

    assume "x  grid p {d}" and "x  p" and "x  grid ?r {d}"
    hence "lv p d  lv x d" using grid_single_level and d < length p by auto
    moreover have "lv p d  lv x d"
    proof (rule ccontr)
      assume "¬ lv p d  lv x d"
      hence "level x = level p" using x  grid p {d} and grid_level_eq[where ds="{d}"] by auto
      hence "x = p" using grid_Start and x  grid p {d} by auto
      thus False using x  p by auto
    qed
    ultimately have "lv p d < lv x d" by auto
    hence "lv ?r d  lv x d" and "?r  grid p {d}" using child_lv and d < length p by auto
    with d < length p and x  grid p {d}
    have r_range: "¬ ?nr_r ?r  ¬ ?nr_l ?r  x  grid ?r {d}"
      using grid_part[where p="?r" and p'=x and b=p and d=d] by auto
    have "x  grid ?r {d}  ?nr_l ?r  ?nr_r ?r" by (rule ccontr, auto simp add: r_range)
    hence "?nr_l ?r  ?nr_r ?r" using x  grid ?r {d} by auto

    have gt0: "lv x d - lv p d > 0" using lv p d < lv x d by auto

    have ix_shift: "ix ?r d = ix ?l d + 2" and lv_lr: "lv ?r d = lv ?l d" and right1: "!! x :: int. x + 2 - 1 = x + 1"
      using d < length p and child_ix and child_lv by auto

    have "lv x d - lv p d = Suc (lv x d - (lv p d + 1))"
      using gt0 by auto
    hence lv_shift: "!! y :: int. y * 2 ^ (lv x d - lv p d) = y * 2 * 2 ^ (lv x d - (lv p d + 1))"
      by auto

    have "ix x d < (ix p d + 1) * 2 ^ (lv x d - lv p d)"
      using x  grid p {d} grid_estimate and d < length p by auto
    also have " = (ix ?r d + 1) * 2 ^ (lv x d - lv ?r d)"
      using lv p d < lv x d and ix_r_eq and lv_r_eq lv_shift[where y="ix p d + 1"] by auto
    finally have "?nr_l ?r" using ?nr_l ?r  ?nr_r ?r by auto
    hence r_bound: "(ix ?l d + 1) * 2 ^ (lv x d - lv ?l d) > ix x d"
      unfolding ix_shift lv_lr using right1 by auto

    have "(ix ?l d - 1) * 2 ^ (lv x d - lv ?l d) = (ix p d - 1) * 2 * 2 ^ (lv x d - (lv p d + 1))"
      unfolding ix_l_eq lv_l_eq by auto
    also have " = (ix p d - 1) * 2 ^ (lv x d - lv p d)"
      using lv_shift[where y="ix p d - 1"] by auto
    also have "  < ix x d"
      using x  grid p {d} grid_estimate and d < length p by auto
    finally have l_bound: "(ix ?l d - 1) * 2 ^ (lv x d - lv ?l d) < ix x d" .

    from l_bound r_bound d < length p and x  grid p {d} lv ?r d  lv x d and lv_lr
    show "x  grid ?l {d}" using grid_part[where p="?l" and p'=x and d=d] by auto
  qed (auto simp add: child_def)
  thus ?thesis by (auto intro: grid_child)
qed
lemma grid_change_dim: assumes grid: "p  grid b ds"
  shows "p[d := X]  grid (b[d := X]) ds"
  using grid
proof induct
  case (Child p d' dir)
  show ?case
  proof (cases "d  d'")
    case True
    have "(child p dir d')[d := X] = child (p[d := X]) dir d'"
      unfolding child_def and ix_def and lv_def
      unfolding list_update_swap[OF d  d'] and nth_list_update_neq[OF d  d'] ..
    thus ?thesis using Child by auto
  next
    case False hence "d = d'" by auto
    with Child show ?thesis unfolding child_def d = d' list_update_overwrite by auto
  qed
qed auto
lemma grid_change_dim_child: assumes grid: "p  grid b ds" and "d  ds"
  shows "child p dir d  grid (child b dir d) ds"
proof (cases "d < length b")
  case True thus ?thesis using grid_change_dim[OF grid]
    unfolding child_def lv_def ix_def grid_invariant[OF True d  ds grid] by auto
next
  case False hence "length b  d" and "length p  d" using grid by auto
  thus ?thesis unfolding child_def using list_update_beyond assms by auto
qed
lemma grid_split: assumes grid: "p  grid b (ds'  ds)" shows " x  grid b ds. p  grid x ds'"
  using grid
proof induct
  case (Child p d dir)
  show ?case
  proof (cases "d  ds'")
    case True with Child show ?thesis by auto
  next
    case False
    hence "d  ds" using Child by auto
    obtain x where "x  grid b ds" and "p  grid x ds'" using Child by auto
    hence "child x dir d  grid b ds" using d  ds by auto
    moreover have "child p dir d  grid (child x dir d) ds'"
      using p  grid x ds' False and grid_change_dim_child by auto
    ultimately show ?thesis by auto
  qed
qed auto
lemma grid_union_eq: "( p  grid b ds. grid p ds') = grid b (ds'  ds)"
  using grid_split and grid_transitive[where ds''="ds'  ds" and ds=ds' and ds'=ds, OF _ _ Un_upper2 Un_upper1] by auto
lemma grid_onedim_split:
  "grid b (ds  {d}) = grid b ds  grid (child b left d) (ds  {d})  grid (child b right d) (ds  {d})"
  (is "_ = ?g  ?l (ds  {d})  ?r (ds  {d})")
proof -
  have "?g  ?l (ds  {d})  ?r (ds  {d}) = ?g  ( p  ?l {d}. grid p ds)  ( p  ?r {d}. grid p ds)"
    unfolding grid_union_eq ..
  also have " = ( p  ({b}  ?l {d}  ?r {d}). grid p ds)" by auto
  also have " = ( p  grid b {d}. grid p ds)" unfolding grid_partition[where p=b] ..
  finally show ?thesis unfolding grid_union_eq by auto
qed
lemma grid_child_without_parent: assumes grid: "p  grid (child b dir d) ds" (is "p  grid ?c ds") and "d < length b"
  shows "p  b"
proof -
  have "level ?c  level p" using grid by (rule grid_level)
  hence "level b < level p" using child_level and d < length b by auto
  thus ?thesis by auto
qed
lemma grid_disjunct':
  assumes "p  grid b ds" and "p'  grid b ds" and "x  grid p ds'" and "p  p'" and "ds  ds' = {}"
  shows "x  grid p' ds'"
proof (rule ccontr)
  assume "¬ x  grid p' ds'" hence "x  grid p' ds'" by auto
  have l: "length b = length p" and l': "length b = length p'" using p  grid b ds and p'  grid b ds by auto
  hence "length p' = length p" by auto
  moreover have " d < length p'. p' ! d = p ! d"
  proof (rule allI, rule impI)
    fix d assume dl': "d < length p'" hence "d < length b" using l' by auto
    hence dl: "d < length p" using l by auto
    show "p' ! d = p ! d"
    proof (cases "d  ds'")
      case True with ds  ds' = {} have "d  ds" by auto
      hence "p' ! d = b ! d" and "p ! d = b ! d"
        using d < length b p'  grid b ds and p  grid b ds and grid_invariant by auto
      thus ?thesis by auto
    next
      case False
      show ?thesis
        using grid_invariant[OF dl' False x  grid p' ds']
          and grid_invariant[OF dl False x  grid p ds'] by auto
    qed
  qed
  ultimately have "p' = p" by (metis nth_equalityI)
  thus False using p  p' by auto
qed
lemma grid_split1: assumes grid: "p  grid b (ds'  ds)" and "ds  ds' = {}"
  shows "∃! x  grid b ds. p  grid x ds'"
proof (rule ex_ex1I)
  obtain x where "x  grid b ds" and "p  grid x ds'" using grid_split[OF grid] by auto
  thus " x. x  grid b ds  p  grid x ds'" by auto
next
  fix x y
  assume "x  grid b ds  p  grid x ds'" and "y  grid b ds  p  grid y ds'"
  hence "x  grid b ds" and "p  grid x ds'" and "y  grid b ds" and "p  grid y ds'" by auto
  show "x = y"
  proof (rule ccontr)
    assume "x  y"
    from grid_disjunct'[OF x  grid b ds y  grid b ds p  grid x ds' this ds  ds' = {}]
    show False using p  grid y ds' by auto
  qed
qed

subsection ‹ Grid Restricted to a Level ›

definition lgrid :: "grid_point  nat set  nat  grid_point set"
where "lgrid b ds lm = { p  grid b ds. level p < lm }"

lemma lgridI[intro]:
  " p  grid b ds ; level p < lm   p  lgrid b ds lm"
  unfolding lgrid_def by simp

lemma lgridE[elim]:
  assumes "p  lgrid b ds lm"
  assumes " p  grid b ds ; level p < lm   P"
  shows P
  using assms unfolding lgrid_def by auto

lemma lgridI_child[intro]:
  "d  ds  p  lgrid (child b dir d) ds lm  p  lgrid b ds lm"
  by (auto intro: grid_child)

lemma lgrid_empty[simp]: "lgrid p ds (level p) = {}"
proof (rule equals0I)
  fix p' assume "p'  lgrid p ds (level p)"
  hence "level p' < level p" and "level p  level p'" by auto
  thus False by auto
qed

lemma lgrid_empty': assumes "lm  level p" shows "lgrid p ds lm = {}"
proof (rule equals0I)
  fix p' assume "p'  lgrid p ds lm"
  hence "level p' < lm" and "level p  level p'" by auto
  thus False using lm  level p by auto
qed

lemma grid_not_child:
  assumes [simp]: "d < length p"
  shows "p  grid (child p dir d) ds"
proof (rule ccontr)
  assume "¬ ?thesis"
  have "level p < level (child p dir d)" by auto
  with grid_level[OF ¬ ?thesis[unfolded not_not]]
  show False by auto
qed

subsection ‹ Unbounded Sparse Grid ›

definition sparsegrid' :: "nat  grid_point set"
where
  "sparsegrid' dm = grid (start dm) { 0 ..< dm }"

lemma grid_subset_alldim:
  assumes p: "p  grid b ds"
  defines "dm  length b"
  shows "p  grid b {0..<dm}"
proof -
  have "ds  {dm..}  ds  {0..<dm} = ds" by auto
  from gridgen_dim_restrict[where ds="ds  {0..<dm}" and ds'="ds  {dm..}"] this
  have "ds  {0..<dm}  {0..<dm}"
    and "p  grid b (ds  {0..<dm})" using p unfolding dm_def by auto
  thus ?thesis by (rule grid_union_dims)
qed

lemma sparsegrid'_length[simp]:
  "b  sparsegrid' dm  length b = dm" unfolding sparsegrid'_def by auto

lemma sparsegrid'I[intro]:
  assumes b: "b  sparsegrid' dm" and p: "p  grid b ds"
  shows "p  sparsegrid' dm"
  using sparsegrid'_length[OF b] b
       grid_transitive[OF grid_subset_alldim[OF p], where c="start dm" and ds''="{0..<dm}"]
  unfolding sparsegrid'_def by auto

lemma sparsegrid'_start:
  assumes "b  grid (start dm) ds"
  shows "b  sparsegrid' dm"
  unfolding sparsegrid'_def
  using grid_subset_alldim[OF assms] by simp

subsection ‹ Sparse Grid ›

definition sparsegrid :: "nat  nat  grid_point set"
where
  "sparsegrid dm lm = lgrid (start dm) { 0 ..< dm } lm"

lemma sparsegrid_length: "p  sparsegrid dm lm  length p = dm"
  by (auto simp: sparsegrid_def)

lemma sparsegrid_subset[intro]: "p  sparsegrid dm lm  p  sparsegrid' dm"
  unfolding sparsegrid_def sparsegrid'_def lgrid_def by auto

lemma sparsegridI[intro]:
  assumes "p  sparsegrid' dm" and "level p < lm"
  shows "p  sparsegrid dm lm"
  using assms unfolding sparsegrid'_def sparsegrid_def lgrid_def by auto

lemma sparsegrid_start:
  assumes "b  lgrid (start dm) ds lm"
  shows "b  sparsegrid dm lm"
proof
  have "b  grid (start dm) ds" using assms by auto
  thus "b  sparsegrid' dm" by (rule sparsegrid'_start)
qed (insert assms, auto)

lemma sparsegridE[elim]:
  assumes "p  sparsegrid dm lm"
  shows "p  sparsegrid' dm" and "level p < lm"
  using assms unfolding sparsegrid'_def sparsegrid_def lgrid_def by auto

subsection ‹ Compute Sparse Grid Points ›

fun gridgen :: "grid_point  nat set  nat  grid_point list"
where
  "gridgen p ds 0 = []"
| "gridgen p ds (Suc l) = (let
      sub = λ d. gridgen (child p left d) { d'  ds . d'  d } l @
                 gridgen (child p right d) { d'  ds . d'  d } l
      in p # concat (map sub [ d  [0 ..< length p]. d  ds]))"

lemma gridgen_lgrid_eq: "set (gridgen p ds l) = lgrid p ds (level p + l)"
proof (induct l arbitrary: p ds)
  case (Suc l)
  let "?subg dir d" = "set (gridgen (child p dir d) { d'  ds . d'   d } l)"
  let "?sub dir d" = "lgrid (child p dir d) { d'  ds . d'   d } (level p + Suc l)"
  let "?union F dm" = "{p}  ( d  { d  ds. d < dm }. F left d  F right d)"

  have hyp: "!! dir d. d < length p  ?subg dir d = ?sub dir d"
    using Suc.hyps using child_level by auto

  { fix dm assume "dm  length p"
    hence "?union ?sub dm = lgrid p {d  ds. d < dm} (level p + Suc l)"
    proof (induct dm)
      case (Suc dm)
      hence "dm  length p" by auto

      let ?l = "child p left dm" and ?r = "child p right dm"

      have p_lgrid: "p  lgrid p {d  ds. d < dm} (level p + Suc l)" by auto

      show ?case
      proof (cases "dm  ds")
        case True
        let ?ds = "{d  ds. d < dm}  {dm}"
        have ds_eq: "{d'  ds. d'  dm} = ?ds" using True by auto
        have ds_eq': "{d  ds. d < Suc dm} = {d  ds. d < dm }  {dm}" using True by auto

        have "?union ?sub (Suc dm) = ?union ?sub dm  ({p}  ?sub left dm  ?sub right dm)"
          unfolding ds_eq' by auto
        also have " = lgrid p {d  ds. d < dm} (level p + Suc l)  ?sub left dm  ?sub right dm"
          unfolding Suc.hyps[OF dm  length p] using p_lgrid by auto
        also have " = {p'  grid p {d  ds. d<dm}  (grid ?l ?ds)  (grid ?r ?ds).
          level p' < level p + Suc l}" unfolding lgrid_def ds_eq by auto
        also have " = lgrid p {d  ds. d < Suc dm} (level p + Suc l)"
          unfolding lgrid_def ds_eq' unfolding grid_onedim_split[where b=p] ..
        finally show ?thesis .
      next
        case False hence "{d  ds. d < Suc dm} = {d  ds. d < dm  d = dm}" by auto
        hence ds_eq: "{d  ds. d < Suc dm} = {d  ds. d < dm}" using dm  ds by auto
        show ?thesis unfolding ds_eq Suc.hyps[OF dm  length p] ..
      qed
    next case 0 thus ?case unfolding lgrid_def by auto
    qed }
  hence "?union ?sub (length p) = lgrid p {d  ds. d < length p} (level p + Suc l)" by auto
  hence union_lgrid_eq: "?union ?sub (length p) = lgrid p ds (level p + Suc l)"
    unfolding lgrid_def using grid_dim_remove_outer by auto

  have "set (gridgen p ds (Suc l)) = ?union ?subg (length p)"
    unfolding gridgen.simps and Let_def by auto
  hence "set (gridgen p ds (Suc l)) = ?union ?sub (length p)"
    using hyp by auto
  also have " = lgrid p ds (level p + Suc l)"
    using union_lgrid_eq .
  finally show ?case .
qed auto

lemma gridgen_distinct: "distinct (gridgen p ds l)"
proof (induct l arbitrary: p ds)
  case (Suc l)
  let ?ds = "[d  [0..<length p]. d  ds]"
  let "?left d" = "gridgen (child p left d) { d'  ds . d'  d } l"
  and "?right d" = "gridgen (child p right d) { d'  ds . d'  d } l"
  let "?sub d" = "?left d @ ?right d"

  have "distinct (concat (map ?sub ?ds))"
  proof (cases l)
    case (Suc l')

    have inj_on: "inj_on ?sub (set ?ds)"
    proof (rule inj_onI, rule ccontr)
      fix d d' assume "d  set ?ds" and "d'  set ?ds"
      hence "d < length p" and "d  set ?ds" and "d' < length p" by auto
      assume *: "?sub d = ?sub d'"
      have in_d: "child p left d  set (?sub d)"
        using d  set ?ds Suc
        by (auto simp add: gridgen_lgrid_eq lgrid_def grid_Start)

      have in_d': "child p left d'  set (?sub d')"
        using d  set ?ds Suc
        by (auto simp add: gridgen_lgrid_eq lgrid_def grid_Start)

      { fix p' d assume "d  set ?ds" and "p'  set (?sub d)"
        hence "lv p d < lv p' d"
          using grid_child_level
          by (auto simp add: gridgen_lgrid_eq lgrid_def grid_child_level) }
      note level_less = this

      assume "d  d'"
      show False
      proof (cases "d' < d")
        case True
        with in_d' ?sub d = ?sub d' level_less[OF d  set ?ds]
        have "lv p d < lv (child p left d') d" by simp
        thus False unfolding lv_def
          using child_invariant[OF d < length p, of left d'] d  d'
          by auto
      next
        case False hence "d < d'" using d  d' by auto
        with in_d ?sub d = ?sub d' level_less[OF d'  set ?ds]
        have "lv p d' < lv (child p left d) d'" by simp
        thus False unfolding lv_def
          using child_invariant[OF d' < length p, of left d] d  d'
          by auto
      qed
    qed

    show ?thesis
    proof (rule distinct_concat)
      show "distinct (map ?sub ?ds)"
        unfolding distinct_map using inj_on by simp
    next
      fix ys assume "ys  set (map ?sub ?ds)"
      then obtain d where "d  ds" and "d < length p"
        and *: "ys = ?sub d" by auto

      show "distinct ys" unfolding *
        using grid_disjunct[OF d < length p, of "{d'  ds. d'  d}"]
          gridgen_lgrid_eq lgrid_def distinct (?left d) distinct (?right d)
        by auto
    next
      fix ys zs
      assume "ys  set (map ?sub ?ds)"
      then obtain d where ys: "ys = ?sub d" and "d  set ?ds" by auto
      hence "d < length p" by auto

      assume "zs  set (map ?sub ?ds)"
      then obtain d' where zs: "zs = ?sub d'" and "d'  set ?ds" by auto
      hence "d' < length p" by auto

      assume "ys  zs"
      hence "d'  d" unfolding ys zs by auto

      show "set ys  set zs = {}"
      proof (rule ccontr)
        assume "¬ ?thesis"
        then obtain p' where "p'  set (?sub d)" and "p'  set (?sub d')"
          unfolding ys zs by auto

        hence "lv p d < lv p' d" "lv p d' < lv p' d'"
          using grid_child_level d  set ?ds d'  set ?ds
          by (auto simp add: gridgen_lgrid_eq lgrid_def grid_child_level)

        show False
        proof (cases "d < d'")
          case True
          from p'  set (?sub d)
          have "p ! d' = p' ! d'"
            using grid_invariant[of d' "child p right d" "{d'  ds. d'  d}"]
            using grid_invariant[of d' "child p left d" "{d'  ds. d'  d}"]
            using child_invariant[of d' _ _ d] d < d' d' < length p
            using gridgen_lgrid_eq lgrid_def by auto
          thus False using lv p d' < lv p' d' unfolding lv_def by auto
        next
          case False hence "d' < d" using d'  d by simp
          from p'  set (?sub d')
          have "p ! d = p' ! d"
            using grid_invariant[of d "child p right d'" "{d  ds. d  d'}"]
            using grid_invariant[of d "child p left d'" "{d  ds. d  d'}"]
            using child_invariant[of d _ _ d'] d' < d d < length p
            using gridgen_lgrid_eq lgrid_def by auto
          thus False using lv p d < lv p' d unfolding lv_def by auto
        qed
      qed
    qed
  qed (simp add: map_replicate_const)
  moreover
  have "p  set (concat (map ?sub ?ds))"
    using gridgen_lgrid_eq lgrid_def grid_not_child[of _ p] by simp
  ultimately show ?case
    unfolding gridgen.simps Let_def distinct.simps by simp
qed auto

lemma lgrid_finite: "finite (lgrid b ds lm)"
proof (cases "level b  lm")
  case True from iffD1[OF le_iff_add True]
  obtain l where l: "lm = level b + l" by auto
  show ?thesis unfolding l gridgen_lgrid_eq[symmetric] by auto
next
  case False hence "!! x. x  grid b ds  (¬ level x < lm)"
  proof -
    fix x assume "x  grid b ds"
    from grid_level[OF this] show "¬ level x < lm" using False by auto
  qed
  hence "lgrid b ds lm = {}" unfolding lgrid_def by auto
  thus ?thesis by auto
qed

lemma lgrid_sum:
  fixes F :: "grid_point  real"
  assumes "d < length b" and "level b < lm"
  shows "( p  lgrid b {d} lm. F p) =
          ( p  lgrid (child b left d) {d} lm. F p) + ( p  lgrid (child b right d) {d} lm. F p) + F b"
  (is "( p  ?grid b. F p) = ( p  ?grid ?l . F p) + (?sum (?grid ?r)) + F b")
proof -
  have "!! dir. b  ?grid (child b dir d)"
    using grid_child_without_parent[where ds="{d}"] and d < length b and lgrid_def by auto
  hence b_distinct: "b  (?grid ?l  ?grid ?r)" by auto

  have "?grid ?l  ?grid ?r = {}"
    unfolding lgrid_def using grid_disjunct and d < length b by auto
  from lgrid_finite lgrid_finite and this
  have child_eq: "?sum ((?grid ?l)  (?grid ?r)) = ?sum (?grid ?l) + ?sum (?grid ?r)"
    by (rule sum.union_disjoint)

  have "?grid b = {b}  (?grid ?l)  (?grid ?r)" unfolding lgrid_def grid_partition[where p=b] using assms by auto
  hence "?sum (?grid b) = F b + ?sum ((?grid ?l)  (?grid ?r))" using b_distinct and lgrid_finite by auto
  thus ?thesis using child_eq by auto
qed

subsection ‹ Base Points ›

definition base :: "nat set  grid_point  grid_point"
where "base ds p = (THE b. b  grid (start (length p)) ({0 ..< length p} - ds)  p  grid b ds)"

lemma baseE: assumes p_grid: "p  sparsegrid' dm"
  shows "base ds p  grid (start dm) ({0..<dm} - ds)"
  and "p  grid (base ds p) ds"
proof -
  from p_grid[unfolded sparsegrid'_def]
  have *: "∃! x  grid (start dm) ({0..<dm} - ds). p  grid x ds"
    by (intro grid_split1) (auto intro: grid_union_dims)
  then obtain x where x_eq: "x  grid (start dm) ({0..<dm} - ds)  p  grid x ds"
    by auto
  with * have "base ds p = x" unfolding base_def by auto
  thus "base ds p  grid (start dm) ({0..<dm} - ds)" and "p  grid (base ds p) ds"
    using x_eq by auto
qed

lemma baseI: assumes x_grid: "x  grid (start dm) ({0..<dm} - ds)" and p_xgrid: "p  grid x ds"
  shows "base ds p = x"
proof -
  have "p  grid (start dm) (ds  ({0..<dm} - ds))"
    using grid_transitive[OF p_xgrid x_grid, where ds''="ds  ({0..<dm} - ds)"] by auto
  moreover have "ds  ({0..<dm} - ds) = {}" by auto
  ultimately have "∃! x  grid (start dm) ({0..<dm} - ds). p  grid x ds"
    using grid_split1[where p=p and b="start dm" and ds'=ds and ds="{0..<dm} - ds"] by auto
  thus "base ds p = x" using x_grid p_xgrid unfolding base_def by auto
qed

lemma base_empty: assumes p_grid: "p  sparsegrid' dm" shows "base {} p = p"
  using grid_empty_ds and p_grid and grid_split1[where ds="{0..<dm}" and ds'="{}"] unfolding base_def sparsegrid'_def by auto

lemma base_start_eq: assumes p_spg: "p  sparsegrid dm lm"
  shows "start dm = base {0..<dm} p"
proof -
  from p_spg
  have "start dm  grid (start dm) ({0..<dm} - {0..<dm})"
    and "p  grid (start dm) {0..<dm}" using sparsegrid'_def by auto
  from baseI[OF this(1) this(2)] show ?thesis by auto
qed

lemma base_in_grid: assumes p_grid: "p  sparsegrid' dm" shows "base ds p  grid (start dm) {0..<dm}"
proof -
  let ?ds = "ds  {0..<dm}"
  have ds_eq: "{ d  ?ds. d < length (start dm) } = { 0..< dm}"
    unfolding start_def by auto
  have "base ds p  grid (start dm) ?ds"
    using grid_union_dims[OF _ baseE(1)[OF p_grid, where ds=ds], where ds'="?ds"] by auto
  thus ?thesis using grid_dim_remove_outer[where b="start dm" and ds="?ds"] unfolding ds_eq by auto
qed

lemma base_grid: assumes p_grid: "p  sparsegrid' dm" shows "grid (base ds p) ds  sparsegrid' dm"
proof
  fix x assume xgrid: "x  grid (base ds p) ds"
  have ds_eq: "{ d  {0..<dm}  ds. d < length (start dm) } = {0..<dm}" by auto
  from grid_transitive[OF xgrid base_in_grid[OF p_grid], where ds''="{0..<dm}  ds"]
  show "x  sparsegrid' dm" unfolding sparsegrid'_def
    using grid_dim_remove_outer[where b="start dm" and ds="{0..<dm}  ds"] unfolding ds_eq unfolding Un_ac(3)[of "{0..<dm}"]
    by auto
qed
lemma base_length[simp]: assumes p_grid: "p  sparsegrid' dm" shows "length (base ds p) = dm"
proof -
  from baseE[OF p_grid] have "base ds p  grid (start dm) ({0..<dm} - ds)" by auto
  thus ?thesis by auto
qed
lemma base_in[simp]: assumes "d < dm" and "d  ds" and p_grid: "p  sparsegrid' dm" shows "base ds p ! d = start dm ! d"
proof -
  have ds: "d  {0..<dm} - ds" using d  ds by auto
  have "d < length (start dm)" using d < dm by auto
  with grid_invariant[OF this ds] baseE(1)[OF p_grid] show ?thesis by auto
qed
lemma base_out[simp]: assumes "d < dm" and "d  ds" and p_grid: "p  sparsegrid' dm" shows "base ds p ! d = p ! d"
proof -
  have "d < length (base ds p)" using base_length[OF p_grid] d < dm by auto
  with grid_invariant[OF this d  ds] baseE(2)[OF p_grid] show ?thesis by auto
qed
lemma base_base: assumes p_grid: "p  sparsegrid' dm" shows "base ds (base ds' p) = base (ds  ds') p"
proof (rule nth_equalityI)
  have b_spg: "base ds' p  sparsegrid' dm" unfolding sparsegrid'_def
    using grid_union_dims[OF Diff_subset[where A="{0..<dm}" and B="ds'"] baseE(1)[OF p_grid]] .
  from base_length[OF b_spg] base_length[OF p_grid] show "length (base ds (base ds' p)) = length (base (ds  ds') p)" by auto

  show "base ds (base ds' p) ! i = base (ds  ds') p ! i" if "i < length (base ds (base ds' p))" for i
  proof -
    have "i < dm" using that base_length[OF b_spg] by auto
    show "base ds (base ds' p) ! i = base (ds  ds') p ! i"
    proof (cases "i  ds  ds'")
      case True
      show ?thesis
      proof (cases "i  ds")
        case True from base_in[OF i < dm i  ds  ds' p_grid] base_in[OF i < dm this b_spg] show ?thesis by auto
      next
        case False hence "i  ds'" using i  ds  ds' by auto
        from base_in[OF i < dm i  ds  ds' p_grid] base_out[OF i < dm i  ds b_spg] base_in[OF i < dm i  ds' p_grid] show ?thesis by auto
      qed
    next
      case False hence "i  ds" and "i  ds'" by auto
      from base_out[OF i < dm i  ds  ds' p_grid] base_out[OF i < dm i  ds b_spg] base_out[OF i < dm i  ds' p_grid] show ?thesis by auto
    qed
  qed
qed
lemma grid_base_out: assumes "d < dm" and "d  ds" and p_grid: "b  sparsegrid' dm" and "p  grid (base ds b) ds"
  shows "p ! d = b ! d"
proof -
  have "base ds b ! d = b ! d" using assms by auto
  moreover have "d < length (base ds b)" using assms by auto
  from grid_invariant[OF this]
  have "p ! d = base ds b ! d" using assms by auto
  ultimately show ?thesis by auto
qed

lemma grid_grid_inj_on: assumes "ds  ds' = {}" shows "inj_on snd (p'grid b ds. p''grid p' ds'. {(p', p'')})"
proof (rule inj_onI)
  fix x y
  assume "x  (p'grid b ds. p''grid p' ds'. {(p', p'')})"
  hence "snd x  grid (fst x) ds'" and "fst x  grid b ds" by auto

  assume "y  (p'grid b ds. p''grid p' ds'. {(p', p'')})"
  hence "snd y  grid (fst y) ds'" and "fst y  grid b ds" by auto

  assume "snd x = snd y"
  have "fst x = fst y"
  proof (rule ccontr)
    assume "fst x  fst y"
    from grid_disjunct'[OF fst x  grid b ds fst y  grid b ds snd x  grid (fst x) ds' this ds  ds' = {}]
    show False using snd y  grid (fst y) ds' unfolding snd x = snd y by auto
  qed
  show "x = y" using prod_eqI[OF fst x = fst y snd x = snd y] .
qed

lemma grid_level_d: assumes "d < length b" and p_grid: "p  grid b {d}" and "p  b" shows "lv p d > lv b d"
proof -
  from p_grid[unfolded grid_partition[where p=b]]
  show ?thesis using grid_child_level using assms by auto
qed

lemma grid_base_base: assumes "b  sparsegrid' dm"
  shows "base ds' b  grid (base ds (base ds' b)) (ds  ds')"
proof -
  from base_grid[OF b  sparsegrid' dm] have "base ds' b  sparsegrid' dm" by auto
  from grid_union_dims[OF _ baseE(2)[OF this], of ds "ds  ds'"] show ?thesis by auto
qed

lemma grid_base_union: assumes b_spg: "b  sparsegrid' dm" and p_grid: "p  grid (base ds b) ds" and x_grid: "x  grid (base ds' p) ds'"
  shows "x  grid (base (ds  ds') b) (ds  ds')"
proof -
  have ds_union: "ds  ds' = ds'  (ds  ds')" by auto

  from base_grid[OF b_spg] p_grid have p_spg: "p  sparsegrid' dm"  by auto
  with assms and grid_base_base have base_b': "base ds' p  grid (base ds (base ds' p)) (ds  ds')" by auto
  moreover have "base ds' (base ds b) = base ds' (base ds p)" (is "?b = ?p")
  proof (rule nth_equalityI)
    have bb_spg: "base ds b  sparsegrid' dm" using base_grid[OF b_spg] grid.Start by auto
    hence "dm = length (base ds b)" by auto
    have bp_spg: "base ds p  sparsegrid' dm" using base_grid[OF p_spg] grid.Start by auto

    show "length ?b = length ?p" using base_length[OF bp_spg] base_length[OF bb_spg] by auto
    show "?b ! i = ?p ! i" if "i < length ?b" for i
    proof -
      have "i < dm" and "i < length (base ds b)" using that base_length[OF bb_spg] dm = length (base ds b) by auto
      show "?b ! i = ?p ! i"
      proof (cases "i  ds  ds'")
        case True
        hence "!! x. base ds x  sparsegrid' dm  x  sparsegrid' dm  base ds' (base ds x) ! i = (start dm) ! i"
        proof - fix x assume x_spg: "x  sparsegrid' dm" and xb_spg: "base ds x  sparsegrid' dm"
          show "base ds' (base ds x) ! i = (start dm) ! i"
          proof (cases "i  ds'")
            case True from base_in[OF i < dm this xb_spg] show ?thesis .
          next
            case False hence "i  ds" using i  ds  ds' by auto
            from base_out[OF i < dm False xb_spg] base_in[OF i < dm this x_spg] show ?thesis by auto
          qed
        qed
        from this[OF bp_spg p_spg] this[OF bb_spg b_spg] show ?thesis by auto
      next
        case False hence "i  ds" and "i  ds'" by auto
        from grid_invariant[OF i < length (base ds b) i  ds p_grid]
          base_out[OF i < dm i  ds' bp_spg] base_out[OF i < dm i  ds p_spg] base_out[OF i < dm i  ds' bb_spg]
        show ?thesis by auto
      qed
    qed
  qed
  ultimately have "base ds' p  grid (base (ds  ds') b) (ds  ds')"
    by (simp only: base_base[OF p_spg] base_base[OF b_spg] Un_ac(3))
  from grid_transitive[OF x_grid this] show ?thesis using ds_union by auto
qed
lemma grid_base_dim_add: assumes "ds'  ds" and b_spg: "b  sparsegrid' dm" and p_grid: "p  grid (base ds' b) ds'"
  shows "p  grid (base ds b) ds"
proof -
  have ds_eq: "ds'  ds = ds" using assms by auto

  have "p  sparsegrid' dm" using base_grid[OF b_spg] p_grid by auto
  hence "p  grid (base ds p) ds" using baseE by auto
  from grid_base_union[OF b_spg p_grid this]
  show ?thesis using ds_eq by auto
qed
lemma grid_replace_dim: assumes "d < length b'" and "d < length b" and p_grid: "p  grid b ds" and p'_grid: "p'  grid b' ds"
  shows "p[d := p' ! d]  grid (b[d := b' ! d]) ds" (is "_  grid ?b ds")
  using p'_grid and p_grid
proof induct
  case (Child p'' d' dir)
  hence p''_grid: "p[d := p'' ! d]  grid ?b ds" and "d < length p''" using assms by auto
  have "d < length p" using p_grid assms by auto
  thus ?case
  proof (cases "d' = d")
    case True
    from grid.Child[OF p''_grid d'  ds]
    show ?thesis unfolding child_def ix_def lv_def list_update_overwrite d' = d nth_list_update_eq[OF d < length p''] nth_list_update_eq[OF d < length p] .
  next
    case False
    show ?thesis unfolding child_def nth_list_update_neq[OF False] using Child by auto
  qed
qed (rule grid_change_dim)
lemma grid_shift_base:
  assumes ds_dj: "ds  ds' = {}" and b_spg: "b  sparsegrid' dm" and p_grid: "p  grid (base (ds'  ds) b) (ds'  ds)"
  shows "base ds' p  grid (base (ds  ds') b) ds"
proof -
  from grid_split[OF p_grid]
  obtain x where x_grid: "x  grid (base (ds'  ds) b) ds" and p_xgrid: "p  grid x ds'" by auto
  from grid_union_dims[OF _ this(1)]
  have x_spg: "x  sparsegrid' dm" using base_grid[OF b_spg] by auto

  have b_len: "length (base (ds'  ds) b) = dm" using base_length[OF b_spg] by auto

  define d' where "d' = dm"
  moreover have "d'  dm  x  grid (start dm) ({0..<dm} - {d  ds'. d < d'})"
  proof (induct d')
    case (Suc d')
    with b_len have d'_b: "d' < length (base (ds'  ds) b)" by auto
    show ?case
    proof (cases "d'  ds'")
      case True hence "d'  ds" and "d'  ds'  ds" using ds_dj by auto
      have "{0..<dm} - {d  ds'. d < d'} = ({0..<dm} - {d  ds'. d < d'}) - {d'}  {d'}" using Suc d'  dm by auto
      also have " = ({0..<dm} - {d  ds'. d < Suc d'})  {d'}" by auto
      finally have x_g: "x  grid (start dm) ({d'}  ({0..<dm} - {d  ds'. d < Suc d'}))" using Suc by auto
      from grid_invariant[OF d'_b d'  ds x_grid] base_in[OF _ d'  ds'  ds b_spg] Suc d'  dm
      have "x ! d' = start dm ! d'" by auto
      from grid_dim_remove[OF x_g this] show ?thesis .
    next
      case False
      hence "{d  ds'. d < Suc d'} = {d  ds'. d < d'  d = d'}" by auto
      also have " = {d  ds'. d < d'}" using False by auto
      finally show ?thesis using Suc by auto
    qed
  next
    case 0 show ?case using x_spg[unfolded sparsegrid'_def] by auto
  qed
  moreover have "{0..<dm} - ds' = {0..<dm} - {d  ds'. d < dm}" by auto
  ultimately have "x  grid (start dm) ({0..<dm} - ds')" by auto
  from baseI[OF this p_xgrid] and x_grid
  show ?thesis by (auto simp: Un_ac(3))
qed

subsection ‹ Lift Operation over all Grid Points ›

definition lift :: "(nat  nat  grid_point  vector  vector)  nat  nat  nat  vector  vector"
where "lift f dm lm d = foldr (λ p. f d (lm - level p) p) (gridgen (start dm) ({ 0 ..< dm } - { d }) lm)"

lemma lift:
  assumes "d < dm" and "p  sparsegrid dm lm"
  and Fintro: " l b p α.  b  lgrid (start dm) ({0..<dm} - {d}) lm ;
                          l + level b = lm ; p  sparsegrid dm lm 
              F d l b α p = (if b = base {d} p
                               then ( p'  lgrid b {d} lm. S (α p') p p')
                               else α p)"
  shows "lift F dm lm d α p = ( p'  lgrid (base {d} p) {d} lm. S (α p') p p')"
        (is "?lift = ?S p α")
proof -
  let ?gridgen = "gridgen (start dm) ({0..<dm} - {d}) lm"
  let "?f p" = "F d (lm - level p) p"

  { fix bs β b
    assume "set bs  set ?gridgen" and "distinct bs" and "p  sparsegrid dm lm"
    hence "foldr ?f bs β p = (if base {d} p  set bs then ?S p β else β p)"
    proof (induct bs arbitrary: p)
      case (Cons b bs)
      hence "b  lgrid (start dm) ({0..<dm} - {d}) lm"
        and "(lm - level b) + level b = lm"
        and b_grid: "b  grid (start dm) ({0..<dm} - {d})"
        using lgrid_def gridgen_lgrid_eq by auto
      note F = Fintro[OF this(1,2) p  sparsegrid dm lm]

      have "b  set bs" using distinct (b#bs) by auto

      show ?case
      proof (cases "base {d} p  set (b#bs)")
        case True note base_in_set = this

        show ?thesis
        proof (cases "b = base {d} p")
          case True
          moreover
          { fix p' assume "p'  lgrid b {d} lm"
            hence "p'  grid b {d}" and "level p' < lm" unfolding lgrid_def by auto
            from grid_transitive[OF this(1) b_grid, of "{0..<dm}"] d < dm
              baseI[OF b_grid p'  grid b {d}] b  set bs
              Cons.prems Cons.hyps[of p'] this(2)
            have "foldr ?f bs β p' = β p'" unfolding sparsegrid_def lgrid_def by auto }
          ultimately show ?thesis
            using F base_in_set by auto
        next
          case False
          with base_in_set have "base {d} p  set bs" by auto
          with Cons.hyps[of p] Cons.prems
          have "foldr ?f bs β p = ?S p β" by auto
          thus ?thesis using F base_in_set False by auto
        qed
      next
        case False
        hence "b  base {d} p" by auto
        from False Cons.hyps[of p] Cons.prems
        have "foldr ?f bs β p = β p" by auto
        thus ?thesis using False F b  base {d} p by auto
      qed
    qed auto
  }
  moreover have "base {d} p  set ?gridgen"
  proof -
    have "p  grid (base {d} p) {d}"
      using p  sparsegrid dm lm[THEN sparsegrid_subset] by (rule baseE)
    from grid_level[OF this] baseE(1)[OF sparsegrid_subset[OF p  sparsegrid dm lm]]
    show ?thesis using p  sparsegrid dm lm
      unfolding gridgen_lgrid_eq sparsegrid'_def lgrid_def sparsegrid_def
      by auto
  qed
  ultimately show ?thesis unfolding lift_def
    using gridgen_distinct p  sparsegrid dm lm by auto
qed

subsection ‹ Parent Points ›

definition parents :: "nat  grid_point  grid_point  grid_point set"
where "parents d b p = { x  grid b {d}. p  grid x {d} }"

lemma parents_split: assumes p_grid: "p  grid (child b dir d) {d}"
  shows "parents d b p = { b }  parents d (child b dir d) p"
proof (intro set_eqI iffI)
  let ?chd = "child b dir d" and ?chid = "child b (inv dir) d"
  fix x assume "x  parents d b p"
  hence "x  grid b {d}" and "p  grid x {d}" unfolding parents_def by auto
  hence x_split: "x  {b}  grid ?chd {d}  grid ?chid {d}" using grid_onedim_split[where ds="{}" and b=b] and grid_empty_ds
    by (cases dir, auto)
  thus "x  {b}  parents d (child b dir d) p"
  proof (cases "x = b")
    case False
    have "d < length b"
    proof (rule ccontr)
      assume "¬ d < length b" hence empty: "{d'  {d}. d' < length b} = {}" by auto
      have "x = b" using x  grid b {d}
        unfolding grid_dim_remove_outer[where ds="{d}" and b=b] empty
        using grid_empty_ds by auto
      thus False using ¬ x = b by auto
    qed
    have "x  grid ?chid {d}"
    proof (rule ccontr)
      assume "¬ x  grid ?chid {d}"
      hence "p  grid ?chid {d}" using grid_transitive[OF p  grid x {d}, where ds'="{d}"]
        by auto
      hence "p  grid ?chd {d}" using grid_disjunct[OF d < length b] by (cases dir, auto)
      thus False using p  grid ?chd {d} ..
    qed
    with False and x_split
    have "x  grid ?chd {d}" by auto
    thus ?thesis unfolding parents_def using p  grid x {d} by auto
  qed auto
next
  let ?chd = "child b dir d" and ?chid = "child b (inv dir) d"
  fix x assume x_in: "x  {b}  parents d ?chd p"
  thus "x  parents d b p"
  proof (cases "x = b")
    case False
    hence "x  parents d ?chd p" using x_in by auto
    thus ?thesis unfolding parents_def using grid_child[where b=b] by auto
  next
    from p_grid have "p  grid b {d}" using grid_child[where b=b] by auto
    case True thus ?thesis unfolding parents_def using p  grid b {d} by auto
  qed
qed

lemma parents_no_parent: assumes "d < length b" shows "b  parents d (child b dir d) p" (is "_  parents _ ?ch _")
proof
  assume "b  parents d ?ch p" hence "b  grid ?ch {d}" unfolding parents_def by auto
  from grid_level[OF this]
  have "level b + 1  level b" unfolding child_level[OF d < length b] .
  thus False by auto
qed

lemma parents_subset_lgrid: "parents d b p  lgrid b {d} (level p + 1)"
proof
  fix x assume "x  parents d b p"
  hence "x  grid b {d}" and "p  grid x {d}" unfolding parents_def by auto
  moreover hence "level x  level p" using grid_level by auto
  hence "level x < level p + 1" by auto
  ultimately show "x  lgrid b {d} (level p + 1)" unfolding lgrid_def by auto
qed

lemma parents_finite: "finite (parents d b p)"
  using finite_subset[OF parents_subset_lgrid lgrid_finite] .

lemma parent_sum: assumes p_grid: "p  grid (child b dir d) {d}" and "d < length b"
  shows "( x  parents d b p. F x) = F b + ( x  parents d (child b dir d) p. F x)"
  unfolding parents_split[OF p_grid] using parents_no_parent[OF d < length b, where dir=dir and p=p] using parents_finite
  by auto

lemma parents_single: "parents d b b = { b }"
proof
  have "parents d b b  lgrid b {d} (level b + (Suc 0))" using parents_subset_lgrid by auto
  also have " = {b}" unfolding gridgen_lgrid_eq[symmetric] gridgen.simps Let_def by auto
  finally show "parents d b b  { b }" .
next
  have "b  parents d b b" unfolding parents_def by auto
  thus "{ b }  parents d b b" by auto
qed

lemma grid_single_dimensional_specification:
  assumes "d < length b"
  and "odd i"
  and "lv b d + l' = l"
  and "i < (ix b d + 1) * 2^l'"
  and "i > (ix b d - 1) * 2^l'"
  shows "b[d := (l,i)]  grid b {d}"
using assms proof (induct l' arbitrary: b)
  case 0
  hence "i = ix b d" and "l = lv b d" by auto
  thus ?case unfolding ix_def lv_def by auto
next
  case (Suc l')

  have "d  {d}" by auto

  show ?case
  proof (rule linorder_cases)
    assume "i = ix b d * 2^(Suc l')"
    hence "even i" by auto
    thus ?thesis using odd i by blast
  next
    assume *: "i < ix b d * 2^(Suc l')"

    let ?b = "child b left d"

    have "d < length ?b" using Suc by auto
    moreover note odd i
    moreover have "lv ?b d + l' = l"
      and "i < (ix ?b d + 1) * 2^l'"
      and "(ix ?b d - 1) * 2^l' < i"
      unfolding child_ix_left[OF Suc.prems(1)]
      using Suc.prems * child_lv by (auto simp add: field_simps)
    ultimately have "?b[d := (l,i)]  grid ?b {d}"
      by (rule Suc.hyps)
    thus ?thesis
      by (auto intro!: grid_child[OF d  {d}, of _ b left]
        simp add: child_def)
  next
    assume *: "ix b d * 2^(Suc l') < i"

    let ?b = "child b right d"

    have "d < length ?b" using Suc by auto
    moreover note odd i
    moreover have "lv ?b d + l' = l"
      and "i < (ix ?b d + 1) * 2^l'"
      and "(ix ?b d - 1) * 2^l' < i"
      unfolding child_ix_right[OF Suc.prems(1)]
      using Suc.prems * child_lv by (auto simp add: field_simps)
    ultimately have "?b[d := (l,i)]  grid ?b {d}"
      by (rule Suc.hyps)
    thus ?thesis
      by (auto intro!: grid_child[OF d  {d}, of _ b right]
        simp add: child_def)
  qed
qed

lemma grid_multi_dimensional_specification:
  assumes "dm  length b" and "length p = length b"
  and " d. d < dm 
    odd (ix p d) 
    lv b d  lv p d 
    ix p d < (ix b d + 1) * 2^(lv p d - lv b d) 
    ix p d > (ix b d - 1) * 2^(lv p d - lv b d)"
    (is " d. d < dm  ?bounded p d")
  and " d.  dm  d ; d < length b   p ! d = b ! d"
  shows "p  grid b {0..<dm}"
using assms proof (induct dm arbitrary: p)
  case 0
  hence "p = b" by (auto intro!: nth_equalityI)
  thus ?case by auto
next
  case (Suc dm)
  hence "dm  length b"
    and "dm < length p" by auto

  let ?p = "p[dm := b ! dm]"

  note dm  length b
  moreover have "length ?p = length b" using length p = length b by simp
  moreover
  {
    fix d assume "d < dm"
    hence *: "d < Suc dm" and "dm  d" by auto
    have "?p ! d = p ! d"
      by (rule nth_list_update_neq[OF dm  d])
    hence "?bounded ?p d"
      using Suc.prems(3)[OF *] lv_def ix_def
      by simp
  }
  moreover
  {
    fix d assume "dm  d" and "d < length b"
    have "?p ! d = b ! d"
    proof (cases "d = dm")
      case True thus ?thesis using d < length b length p = length b by auto
    next
      case False
      hence "Suc dm  d" using dm  d by auto
      thus ?thesis using Suc.prems(4) d < length b by auto
    qed
  }
  ultimately
  have *: "?p  grid b {0..<dm}"
    by (auto intro!: Suc.hyps)

  have "lv b dm  lv p dm" using Suc.prems(3)[OF lessI] by simp

  have [simp]: "lv ?p dm = lv b dm" using lv_def dm < length p by auto
  have [simp]: "ix ?p dm = ix b dm" using ix_def dm < length p by auto
  have [simp]: "p[dm := (lv p dm, ix p dm)] = p"
    using lv_def ix_def dm < length p by auto
  have "dm < length ?p" and
    [simp]: "lv b dm + (lv p dm - lv b dm) = lv p dm"
    using dm < length p lv b dm  lv p dm by auto
  from grid_single_dimensional_specification[OF this(1),
    where l="lv p dm" and i="ix p dm" and l'="lv p dm - lv b dm", simplified]
  have "p  grid ?p {dm}"
    using Suc.prems(3)[OF lessI] by blast
  from grid_transitive[OF this *]
  show ?case by auto
qed

lemma sparsegrid:
  "sparsegrid dm lm = {p.
    length p = dm  level p < lm 
    ( d < dm. odd (ix p d)  0 < ix p d  ix p d < 2^(lv p d + 1))}"
  (is "_ = ?set")
proof (rule equalityI[OF subsetI subsetI])
  fix p
  assume *: "p  sparsegrid dm lm"
  hence "length p = dm" and "level p < lm" unfolding sparsegrid_def by auto
  moreover
  { fix d assume "d < dm"
    hence **: "p  grid (start dm) {0..<dm}" and "d < length (start dm)"
      using * unfolding sparsegrid_def by auto
    have "odd (ix p d)"
    proof (cases "p ! d = start dm ! d")
      case True
      thus ?thesis unfolding start_def using d < dm ix_def by auto
    next
      case False
      from grid_odd[OF _ this **]
      show ?thesis using d < dm by auto
    qed
    hence "odd (ix p d)  0 < ix p d  ix p d < 2^(lv p d + 1)"
      using grid_estimate[OF d < length (start dm) **]
      unfolding ix_def lv_def start_def using d < dm by auto
  }
  ultimately show "p  ?set"
    using sparsegrid_def lgrid_def by auto
next
  fix p
  assume "p  ?set"
  with grid_multi_dimensional_specification[of dm "start dm" p]
  have "p  grid (start dm) {0..<dm}" and "level p < lm"
    by auto
  thus "p  sparsegrid dm lm"
    unfolding sparsegrid_def lgrid_def by auto
qed

end