thys2/ClosedFormsBounds.thy
author Chengsong
Tue, 15 Mar 2022 16:37:41 +0000
changeset 451 7a016eeb118d
parent 450 dabd25e8e4e9
child 452 4aded96b3923
permissions -rw-r--r--
finiteness
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
     1
(**)
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
     2
theory ClosedFormsBounds
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
     3
  imports "GeneralRegexBound" "ClosedForms"
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
     4
begin
449
Chengsong
parents: 448
diff changeset
     5
lemma alts_ders_lambda_shape_ders:
Chengsong
parents: 448
diff changeset
     6
  shows "\<forall>r \<in> set (map (\<lambda>r. rders_simp r ( s)) rs ). \<exists>r1 \<in> set rs. r = rders_simp r1 s"
Chengsong
parents: 448
diff changeset
     7
  by (simp add: image_iff)
Chengsong
parents: 448
diff changeset
     8
Chengsong
parents: 448
diff changeset
     9
lemma rlist_bound:
Chengsong
parents: 448
diff changeset
    10
  shows "\<forall>r \<in> set rs. rsize r \<le> N \<Longrightarrow> sum_list (map rsize rs) \<le> N * (length rs)"
Chengsong
parents: 448
diff changeset
    11
  apply(induct rs)
Chengsong
parents: 448
diff changeset
    12
  apply simp
Chengsong
parents: 448
diff changeset
    13
  by simp
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
    14
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
    15
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
    16
lemma alts_closed_form_bounded: shows
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
    17
"\<forall>r \<in> set rs. \<forall>s. rsize(rders_simp r s ) \<le> N \<Longrightarrow> 
449
Chengsong
parents: 448
diff changeset
    18
rsize (rders_simp (RALTS rs ) s) \<le> max (Suc ( N * (length rs))) (rsize (RALTS rs) )"
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
    19
  apply(induct s)
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
    20
  apply simp
449
Chengsong
parents: 448
diff changeset
    21
  apply(subst alts_closed_form_variant)
Chengsong
parents: 448
diff changeset
    22
   apply force
Chengsong
parents: 448
diff changeset
    23
  apply(subgoal_tac "rsize (rsimp (RALTS (map (\<lambda>r. rders_simp r (a # s)) rs))) \<le> rsize ( (RALTS (map (\<lambda>r. rders_simp r (a # s)) rs)))")
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
    24
   prefer 2
449
Chengsong
parents: 448
diff changeset
    25
  using rsimp_mono apply presburger
Chengsong
parents: 448
diff changeset
    26
  apply(subgoal_tac "rsize (RALTS (map (\<lambda>r. rders_simp r (a # s)) rs)) =
Chengsong
parents: 448
diff changeset
    27
                     Suc (sum_list  (map rsize (map (\<lambda>r. rders_simp r (a # s)) rs)))")
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
    28
  prefer 2
449
Chengsong
parents: 448
diff changeset
    29
  using rsize.simps(4) apply blast
Chengsong
parents: 448
diff changeset
    30
  apply(subgoal_tac "sum_list (map rsize (map (\<lambda>r. rders_simp r (a # s)) rs )) \<le> N *  (length rs) ")
Chengsong
parents: 448
diff changeset
    31
   apply linarith
Chengsong
parents: 448
diff changeset
    32
  apply(subgoal_tac "\<forall>r \<in> set (map (\<lambda>r. rders_simp r (a # s)) rs ). rsize r \<le> N")
Chengsong
parents: 448
diff changeset
    33
  prefer 2
Chengsong
parents: 448
diff changeset
    34
  apply(subgoal_tac "\<forall>r \<in> set (map (\<lambda>r. rders_simp r (a # s)) rs ). \<exists>r1 \<in> set rs. r = rders_simp r1 (a # s)")
Chengsong
parents: 448
diff changeset
    35
  prefer 2
Chengsong
parents: 448
diff changeset
    36
  using alts_ders_lambda_shape_ders apply presburger
Chengsong
parents: 448
diff changeset
    37
   apply metis
Chengsong
parents: 448
diff changeset
    38
  apply(frule rlist_bound)
Chengsong
parents: 448
diff changeset
    39
  by fastforce
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
    40
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
    41
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
    42
lemma alts_simp_ineq_unfold:
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
    43
  shows "rsize (rsimp (RALTS rs)) \<le> Suc (sum_list (map rsize (rdistinct (rflts (map rsimp rs)) {})))"
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
    44
  using rsimp_aalts_smaller by auto
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
    45
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    46
lemma no_elem_distinct :
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    47
  shows "a \<notin> set rs \<Longrightarrow> rdistinct rs rset = rdistinct rs (insert a rset)"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    48
  apply(induct rs arbitrary: rset)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    49
   apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    50
  by force
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    51
  
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    52
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    53
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    54
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    55
inductive good :: "rrexp \<Rightarrow> bool"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    56
  where
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    57
"RZERO \<notin> set rs \<Longrightarrow> good (RALTS rs)"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    58
|"good RZERO"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    59
|"good RONE"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    60
|"good (RCHAR c)"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    61
|"good (RSEQ r1 r2)"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    62
|"good (RSTAR r0)"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    63
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    64
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    65
lemma after_flts_no0:
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    66
  shows "\<forall>r \<in> set rs. good r \<Longrightarrow> RZERO \<notin> set (rflts rs)"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    67
  apply(induct rs)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    68
   apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    69
  apply(case_tac a)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    70
       apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    71
  apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    72
  apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    73
    apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    74
  apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    75
  using good.simps apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    76
  apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    77
  done
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    78
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
    79
lemma flts_has_no_zero:
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    80
  shows "\<forall>r \<in> set rs. good r \<Longrightarrow> rdistinct (rflts rs) rset = rdistinct (rflts rs) (insert RZERO rset)"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    81
  apply(subgoal_tac "RZERO \<notin> set (rflts rs)")  
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    82
   apply (meson no_elem_distinct)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    83
  apply(insert after_flts_no0)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    84
  by blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    85
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    86
lemma shape_of_alts:
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    87
  shows "\<forall>a \<in> alts_set. (\<exists>xs. a = RALTS xs) \<Longrightarrow> RZERO \<notin> alts_set \<and> RONE \<notin> alts_set"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    88
  by fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    89
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    90
lemma quantified2_implies1:
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    91
  shows "\<forall>r. P \<and> Q \<Longrightarrow> \<forall>r. P"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    92
  by auto
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    93
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    94
lemma quantified_quantifiedAB_A:
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    95
  shows " (\<forall>a\<in>alts_set. \<exists>xs. a = RALTS xs \<and> set xs \<subseteq> corr_set) \<Longrightarrow> \<forall>a \<in> alts_set. \<exists>xs. a = RALTS xs"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    96
  by fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    97
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    98
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
    99
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   100
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   101
lemma list_distinct_removal:
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   102
  shows "set rs \<subseteq> rset \<Longrightarrow> rdistinct (rs @ rs1) rset = rdistinct rs1 rset"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   103
  apply(induct rs arbitrary: rset)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   104
   apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   105
  by simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   106
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   107
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   108
lemma rdistinct_mono_list:
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   109
  shows " sum_list (map rsize (rdistinct (x5 @ rs) rset )) \<le>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   110
        ( sum_list (map rsize x5)) + (sum_list (map rsize (rdistinct  rs ((set x5 ) \<union> rset))))"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   111
  apply(induct x5 arbitrary: rs rset)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   112
   apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   113
  apply(case_tac "a \<in> rset")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   114
   apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   115
   apply (simp add: add.assoc insert_absorb trans_le_add2)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   116
  apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   117
  by (metis Un_insert_right)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   118
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   119
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   120
lemma flts_size_reduction_alts:
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   121
  shows " \<And>a rs noalts_set alts_set corr_set x5.
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   122
       \<lbrakk>\<And>noalts_set alts_set corr_set.
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   123
           (\<forall>r\<in>noalts_set. \<forall>xs. r \<noteq> RALTS xs) \<and>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   124
           (\<forall>a\<in>alts_set. \<exists>xs. a = RALTS xs \<and> set xs \<subseteq> corr_set) \<Longrightarrow>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   125
           Suc (sum_list (map rsize (rdistinct (rflts rs) (noalts_set \<union> corr_set))))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   126
           \<le> Suc (sum_list (map rsize (rdistinct rs (insert RZERO (noalts_set \<union> alts_set)))));
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   127
        (\<forall>r\<in>noalts_set. \<forall>xs. r \<noteq> RALTS xs) \<and>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   128
        (\<forall>a\<in>alts_set. \<exists>xs. a = RALTS xs \<and> set xs \<subseteq> corr_set);
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   129
        a = RALTS x5\<rbrakk>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   130
       \<Longrightarrow> Suc (sum_list (map rsize (rdistinct (rflts (a # rs)) (noalts_set \<union> corr_set))))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   131
           \<le> Suc (sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   132
                    (map rsize (rdistinct (a # rs) (insert RZERO (noalts_set \<union> alts_set)))))"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   133
  apply(case_tac "a \<in> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   134
   apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   135
   apply(subgoal_tac "set x5 \<subseteq> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   136
  apply(subst list_distinct_removal)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   137
  apply auto[1]
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   138
    apply presburger
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   139
   apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   140
  apply (subgoal_tac "a \<notin> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   141
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   142
  apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   143
  apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   144
  apply(subgoal_tac "sum_list (map rsize (rdistinct (x5 @ rflts rs) (noalts_set \<union> corr_set))) 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   145
                   \<le> sum_list (map rsize x5 ) + sum_list (map rsize (rdistinct (rflts rs) ((set x5) \<union> (noalts_set \<union> corr_set))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   146
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   147
  using rdistinct_mono_list apply presburger
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   148
  apply(subgoal_tac "insert (RALTS x5) (noalts_set \<union> alts_set) = noalts_set \<union> (insert (RALTS x5) alts_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   149
   apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   150
  apply(subgoal_tac "sum_list (map rsize x5) + 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   151
sum_list (map rsize (rdistinct (rflts rs) (noalts_set \<union> (corr_set \<union> (set x5))))) \<le>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   152
Suc (sum_list (map rsize x5) +
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   153
                   sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   154
                    (map rsize
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   155
                      (rdistinct rs (insert RZERO (noalts_set \<union> insert (RALTS x5) alts_set)))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   156
  
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   157
  apply (simp add: Un_left_commute inf_sup_aci(5))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   158
   apply(subgoal_tac "sum_list (map rsize (rdistinct (rflts rs) (noalts_set \<union> (corr_set \<union> set x5))))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   159
\<le> sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   160
                    (map rsize
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   161
                      (rdistinct rs (insert RZERO (noalts_set \<union> insert (RALTS x5) alts_set))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   162
    apply linarith
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   163
   apply(subgoal_tac "\<forall>r \<in>  insert (RALTS x5) alts_set. \<exists>xs1.( r = RALTS xs1 \<and> set xs1 \<subseteq> corr_set \<union> set x5)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   164
    apply presburger
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   165
   apply (meson insert_iff sup.cobounded2 sup.coboundedI1)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   166
  by blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   167
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   168
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   169
449
Chengsong
parents: 448
diff changeset
   170
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   171
lemma flts_vs_nflts1:
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   172
  shows "(\<forall>r \<in> noalts_set. \<forall>xs. r \<noteq> RALTS xs)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   173
 \<and> (\<forall>a \<in> alts_set. (\<exists>xs. a = RALTS xs \<and> set xs \<subseteq> corr_set))  
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   174
\<Longrightarrow>  (sum_list (map rsize (rdistinct ( rflts rs) (noalts_set \<union> corr_set)  )))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   175
         \<le>  (sum_list (map rsize (rdistinct rs (insert RZERO (noalts_set \<union> alts_set) ))))"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   176
    apply(induct rs arbitrary: noalts_set alts_set corr_set)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   177
   apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   178
  apply(case_tac a)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   179
       apply(case_tac "RZERO \<in> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   180
        apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   181
       apply(subgoal_tac "RZERO \<notin> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   182
        apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   183
       apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   184
      apply(case_tac "RONE \<in> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   185
       apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   186
      apply(subgoal_tac "RONE \<notin> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   187
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   188
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   189
      apply(case_tac "RONE \<in> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   190
       apply(subgoal_tac "rflts (a # rs) = RONE # rflts rs")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   191
        apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   192
        apply(subgoal_tac "rdistinct (RONE # rflts rs) (noalts_set \<union> corr_set) = 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   193
                           rdistinct (rflts rs) (noalts_set \<union> corr_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   194
         apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   195
  apply(subgoal_tac "rdistinct (RONE # rs) (insert RZERO (noalts_set \<union> alts_set)) =
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   196
                     RONE # (rdistinct rs (insert RONE (insert RZERO (noalts_set \<union> alts_set)))) ")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   197
          apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   198
  apply(subgoal_tac "rdistinct (rflts rs) (noalts_set \<union> corr_set) = 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   199
                     rdistinct (rflts rs) (insert RONE (noalts_set \<union> corr_set))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   200
  apply (simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   201
  apply(subgoal_tac "insert RONE (noalts_set \<union> corr_set) = (insert RONE noalts_set) \<union> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   202
            apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   203
  apply(subgoal_tac "insert RONE (insert RZERO (noalts_set \<union> alts_set)) = 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   204
                     insert RZERO ((insert RONE noalts_set) \<union> alts_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   205
             apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   206
  apply(subgoal_tac "  (sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   207
                    (map rsize ( rdistinct rs (insert RZERO (insert RONE noalts_set \<union> alts_set)))))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   208
                   \<le>  (sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   209
                    (map rsize (RONE # rdistinct rs (insert RZERO (insert RONE noalts_set \<union> alts_set)))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   210
  apply (smt (verit, best) dual_order.trans insert_iff rrexp.distinct(15))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   211
  apply (metis (no_types, opaque_lifting) add_Suc_right le_add_same_cancel2 list.simps(9) sum_list.Cons zero_le)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   212
            apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   213
           apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   214
  apply (metis Un_iff insert_absorb)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   215
         apply (metis UnE insertE insert_is_Un rdistinct.simps(2) rrexp.distinct(1))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   216
        apply (meson UnCI rdistinct.simps(2))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   217
  using rflts.simps(4) apply presburger
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   218
      apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   219
      apply(subgoal_tac "insert RONE (noalts_set \<union> corr_set) = (insert RONE noalts_set) \<union> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   220
  apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   221
  apply (metis Un_insert_left insertE rrexp.distinct(15))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   222
      apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   223
     apply(case_tac "a \<in> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   224
      apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   225
  apply(subgoal_tac "a \<notin> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   226
      prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   227
      apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   228
  apply(case_tac "a \<in> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   229
      apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set  \<union> corr_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   230
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   231
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   232
      apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   233
      apply(subgoal_tac "  sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   234
               (map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set))))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   235
\<le>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   236
                sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   237
               (map rsize (rdistinct (a # rs) (insert RZERO (noalts_set \<union> alts_set))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   238
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   239
       apply(subgoal_tac 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   240
"sum_list (map rsize (rdistinct (rflts (a # rs)) ((insert a noalts_set) \<union> corr_set)))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   241
\<le>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   242
 sum_list (map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   243
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   244
       apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   245
  apply(subgoal_tac "(insert a (noalts_set \<union> alts_set)) = (insert a noalts_set) \<union> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   246
        apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   247
        apply(subgoal_tac "noalts_set \<union> corr_set = (insert a noalts_set) \<union> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   248
  apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   249
  apply (metis insertE rrexp.distinct(21))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   250
        apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   251
  
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   252
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   253
  apply force
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   254
     apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   255
     apply (metis Un_insert_left insert_iff rrexp.distinct(21))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   256
    apply(case_tac "a \<in> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   257
     apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   258
  apply(subgoal_tac "a \<notin> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   259
      prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   260
      apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   261
  apply(case_tac "a \<in> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   262
      apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set  \<union> corr_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   263
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   264
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   265
      apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   266
      apply(subgoal_tac "  sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   267
               (map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set))))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   268
\<le>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   269
                sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   270
               (map rsize (rdistinct (a # rs) (insert RZERO (noalts_set \<union> alts_set))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   271
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   272
       apply(subgoal_tac 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   273
"sum_list (map rsize (rdistinct (rflts (a # rs)) ((insert a noalts_set) \<union> corr_set)))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   274
\<le>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   275
 sum_list (map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   276
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   277
       apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   278
  apply(subgoal_tac "(insert a (noalts_set \<union> alts_set)) = (insert a noalts_set) \<union> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   279
        apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   280
        apply(subgoal_tac "noalts_set \<union> corr_set = (insert a noalts_set) \<union> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   281
  apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   282
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   283
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   284
  apply (metis insertE rrexp.distinct(25))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   285
  apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   286
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   287
  apply force
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   288
     apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   289
  
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   290
    apply (metis Un_insert_left insertE rrexp.distinct(25))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   291
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   292
  using Suc_le_mono flts_size_reduction_alts apply presburger
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   293
     apply(case_tac "a \<in> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   294
      apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   295
  apply(subgoal_tac "a \<notin> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   296
      prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   297
      apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   298
  apply(case_tac "a \<in> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   299
      apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set  \<union> corr_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   300
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   301
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   302
      apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   303
      apply(subgoal_tac "  sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   304
               (map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set))))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   305
\<le>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   306
                sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   307
               (map rsize (rdistinct (a # rs) (insert RZERO (noalts_set \<union> alts_set))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   308
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   309
       apply(subgoal_tac 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   310
"sum_list (map rsize (rdistinct (rflts (a # rs)) ((insert a noalts_set) \<union> corr_set)))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   311
\<le>
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   312
 sum_list (map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   313
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   314
       apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   315
  apply(subgoal_tac "(insert a (noalts_set \<union> alts_set)) = (insert a noalts_set) \<union> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   316
        apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   317
        apply(subgoal_tac "noalts_set \<union> corr_set = (insert a noalts_set) \<union> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   318
  apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   319
  apply (metis insertE rrexp.distinct(29))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   320
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   321
        apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   322
  
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   323
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   324
  apply force
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   325
     apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   326
  apply (metis Un_insert_left insert_iff rrexp.distinct(29))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   327
  done
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   328
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   329
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   330
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   331
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   332
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   333
lemma flts_vs_nflts:
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   334
  shows "(\<forall>r \<in> noalts_set. \<forall>xs. r \<noteq> RALTS xs)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   335
 \<and> (\<forall>a \<in> alts_set. (\<exists>xs. a = RALTS xs \<and> set xs \<subseteq> corr_set))  
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   336
\<Longrightarrow> Suc (sum_list (map rsize (rdistinct ( rflts rs) (noalts_set \<union> corr_set)  )))
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   337
         \<le> Suc (sum_list (map rsize (rdistinct rs (insert RZERO (noalts_set \<union> alts_set) ))))"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   338
  apply(induct rs arbitrary: noalts_set alts_set corr_set)
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   339
   apply simp
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   340
  apply(case_tac a)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   341
       apply(case_tac "RZERO \<in> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   342
        apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   343
       apply(subgoal_tac "RZERO \<notin> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   344
        apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   345
       apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   346
      apply(case_tac "RONE \<in> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   347
       apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   348
      apply(subgoal_tac "RONE \<notin> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   349
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   350
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   351
      apply(case_tac "RONE \<in> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   352
       apply(subgoal_tac "rflts (a # rs) = RONE # rflts rs")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   353
        apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   354
        apply(subgoal_tac "rdistinct (RONE # rflts rs) (noalts_set \<union> corr_set) = 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   355
                           rdistinct (rflts rs) (noalts_set \<union> corr_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   356
         apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   357
  apply(subgoal_tac "rdistinct (RONE # rs) (insert RZERO (noalts_set \<union> alts_set)) =
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   358
                     RONE # (rdistinct rs (insert RONE (insert RZERO (noalts_set \<union> alts_set)))) ")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   359
          apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   360
  apply(subgoal_tac "rdistinct (rflts rs) (noalts_set \<union> corr_set) = 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   361
                     rdistinct (rflts rs) (insert RONE (noalts_set \<union> corr_set))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   362
  apply (simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   363
  apply(subgoal_tac "insert RONE (noalts_set \<union> corr_set) = (insert RONE noalts_set) \<union> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   364
            apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   365
  apply(subgoal_tac "insert RONE (insert RZERO (noalts_set \<union> alts_set)) = 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   366
                     insert RZERO ((insert RONE noalts_set) \<union> alts_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   367
             apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   368
  apply(subgoal_tac " Suc (sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   369
                    (map rsize ( rdistinct rs (insert RZERO (insert RONE noalts_set \<union> alts_set)))))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   370
                   \<le> Suc (sum_list
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   371
                    (map rsize (RONE # rdistinct rs (insert RZERO (insert RONE noalts_set \<union> alts_set)))))")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   372
  apply (smt (verit, best) dual_order.trans insert_iff rrexp.distinct(15))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   373
  apply (metis (no_types, opaque_lifting) add_Suc_right le_add_same_cancel2 list.simps(9) sum_list.Cons zero_le)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   374
            apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   375
           apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   376
  apply (metis Un_iff insert_absorb)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   377
         apply (metis UnE insertE insert_is_Un rdistinct.simps(2) rrexp.distinct(1))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   378
        apply (meson UnCI rdistinct.simps(2))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   379
  using rflts.simps(4) apply presburger
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   380
      apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   381
      apply(subgoal_tac "insert RONE (noalts_set \<union> corr_set) = (insert RONE noalts_set) \<union> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   382
  apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   383
  apply (metis Un_insert_left insertE rrexp.distinct(15))
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   384
      apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   385
     apply(case_tac "a \<in> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   386
      apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   387
  apply(subgoal_tac "a \<notin> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   388
      prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   389
      apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   390
  apply(case_tac "a \<in> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   391
      apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set  \<union> corr_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   392
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   393
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   394
      apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   395
  apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   396
  apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   397
     apply(case_tac "a \<in> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   398
      apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   399
  apply(subgoal_tac "a \<notin> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   400
      prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   401
      apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   402
  apply(case_tac "a \<in> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   403
      apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set  \<union> corr_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   404
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   405
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   406
      apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   407
  apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   408
  apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   409
   prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   410
     apply(case_tac "a \<in> noalts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   411
      apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   412
  apply(subgoal_tac "a \<notin> alts_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   413
      prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   414
      apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   415
  apply(case_tac "a \<in> corr_set")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   416
      apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set  \<union> corr_set)")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   417
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   418
  apply fastforce
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   419
      apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   420
  apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   421
  apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   422
  using flts_size_reduction_alts apply presburger
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   423
  done
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   424
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   425
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   426
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   427
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   428
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   429
(*  apply(rutac x = "\<lambda>a rs noalts_set alts_set corr_set. set xs ")*)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   430
  
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   431
449
Chengsong
parents: 448
diff changeset
   432
lemma distinct_simp_ineq_general:
450
Chengsong
parents: 449
diff changeset
   433
  shows "rsimp ` no_simp = has_simp \<and> finite no_simp \<Longrightarrow>Suc (sum_list (map rsize (rdistinct (map rsimp rs) has_simp)))
449
Chengsong
parents: 448
diff changeset
   434
    \<le> Suc (sum_list (map rsize (rdistinct rs no_simp)))"
450
Chengsong
parents: 449
diff changeset
   435
  apply(induct rs arbitrary: no_simp has_simp)
Chengsong
parents: 449
diff changeset
   436
   apply simp
Chengsong
parents: 449
diff changeset
   437
  apply(case_tac "a \<in> no_simp")
Chengsong
parents: 449
diff changeset
   438
  apply(subgoal_tac "rsimp a \<in> has_simp")
Chengsong
parents: 449
diff changeset
   439
    apply auto[1]
Chengsong
parents: 449
diff changeset
   440
  apply blast
Chengsong
parents: 449
diff changeset
   441
  apply(case_tac "rsimp a \<in> no_simp")
Chengsong
parents: 449
diff changeset
   442
   apply(subgoal_tac "rsimp a \<in> has_simp")
Chengsong
parents: 449
diff changeset
   443
  prefer 2
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   444
    apply (simp add: rev_image_eqI rsimp_idem)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   445
  apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   446
   apply (metis finite_insert image_insert insert_absorb trans_le_add2)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   447
  apply(case_tac "rsimp a \<notin> has_simp")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   448
  apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   449
   apply (metis add_mono_thms_linordered_semiring(1) finite.insertI image_insert rsimp_mono)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   450
  apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   451
  by (metis finite.insertI image_insert insert_absorb trans_le_add2)
449
Chengsong
parents: 448
diff changeset
   452
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   453
450
Chengsong
parents: 449
diff changeset
   454
lemma not_mentioned_elem_distinct_strong:
Chengsong
parents: 449
diff changeset
   455
  shows "r \<noteq> a \<Longrightarrow> (r \<in> set (rdistinct rs rset)) = (r \<in> set (rdistinct rs (insert a rset)))"
Chengsong
parents: 449
diff changeset
   456
  apply(induct rs arbitrary: rset)
Chengsong
parents: 449
diff changeset
   457
   apply simp
Chengsong
parents: 449
diff changeset
   458
  by force
Chengsong
parents: 449
diff changeset
   459
Chengsong
parents: 449
diff changeset
   460
lemma not_mentioned_elem_distinct:
Chengsong
parents: 449
diff changeset
   461
  shows "r \<noteq> a \<Longrightarrow> (r \<in> set (rdistinct rs {})) = (r \<in> set (rdistinct rs (insert a {})))"
Chengsong
parents: 449
diff changeset
   462
  by (metis not_mentioned_elem_distinct_strong)
Chengsong
parents: 449
diff changeset
   463
Chengsong
parents: 449
diff changeset
   464
Chengsong
parents: 449
diff changeset
   465
Chengsong
parents: 449
diff changeset
   466
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   467
lemma larger_acc_smaller_distinct_res0:
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   468
  shows " ss \<subseteq> SS \<Longrightarrow> sum_list (map rsize (rdistinct rs SS)) \<le> sum_list (map rsize (rdistinct rs ss))"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   469
  apply(induct rs arbitrary: ss SS)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   470
  apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   471
  apply(case_tac "a \<in> ss")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   472
   apply(subgoal_tac "a \<in> SS")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   473
    apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   474
   apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   475
  apply(case_tac "a \<in> SS")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   476
   apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   477
   apply(subgoal_tac "insert a ss \<subseteq> SS")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   478
    apply simp
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   479
  apply (simp add: trans_le_add2)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   480
  apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   481
  apply(simp)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   482
  apply(subgoal_tac "insert a ss \<subseteq> insert a SS")
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   483
   apply blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   484
  by blast
450
Chengsong
parents: 449
diff changeset
   485
Chengsong
parents: 449
diff changeset
   486
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   487
lemma without_flts_ineq:
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   488
  shows " Suc (sum_list (map rsize (rdistinct (rflts rs) {}) )) \<le> 
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   489
         Suc (sum_list (map rsize (rdistinct (    rs  ) {}  )))"
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   490
proof -
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   491
  have " Suc (sum_list (map rsize (rdistinct (rflts rs) {}) )) \<le>  
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   492
         Suc (sum_list (map rsize (rdistinct rs (insert RZERO {}))))"
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   493
    by (metis empty_iff flts_vs_nflts sup_bot_left)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   494
  also have "... \<le>  Suc (sum_list (map rsize (rdistinct rs {})))" 
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   495
    by (simp add: larger_acc_smaller_distinct_res0)
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   496
  finally show ?thesis sledgehammer
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   497
    by blast
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   498
qed
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   499
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   500
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   501
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   502
448
Chengsong
parents: 447
diff changeset
   503
Chengsong
parents: 447
diff changeset
   504
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   505
lemma distinct_simp_ineq:
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   506
  shows "Suc (sum_list (map rsize (rdistinct (map rsimp rs) {})))
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   507
    \<le> Suc (sum_list (map rsize (rdistinct rs {})))"
448
Chengsong
parents: 447
diff changeset
   508
  
Chengsong
parents: 447
diff changeset
   509
  using distinct_simp_ineq_general by blast
Chengsong
parents: 447
diff changeset
   510
  
Chengsong
parents: 447
diff changeset
   511
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   512
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   513
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   514
lemma alts_simp_control:
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   515
  shows "rsize (rsimp (RALTS rs)) \<le> Suc (sum_list (map rsize (rdistinct rs {})))"
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   516
proof -
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   517
  have "rsize (rsimp (RALTS rs)) \<le> Suc (sum_list (map rsize (rdistinct (rflts (map rsimp rs)) {})))"
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   518
    
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   519
    using alts_simp_ineq_unfold by auto
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   520
  then have "\<dots> \<le> Suc (sum_list (map rsize (rdistinct (map rsimp rs) {})))"
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   521
    using without_flts_ineq by blast
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   522
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   523
  show "rsize (rsimp (RALTS rs)) \<le> Suc (sum_list (map rsize (rdistinct rs {})))"
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   524
    by (meson \<open>Suc (sum_list (map rsize (rdistinct (rflts (map rsimp rs)) {}))) \<le> Suc (sum_list (map rsize (rdistinct (map rsimp rs) {})))\<close> \<open>rsize (rsimp (RALTS rs)) \<le> Suc (sum_list (map rsize (rdistinct (rflts (map rsimp rs)) {})))\<close> distinct_simp_ineq order_trans)
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   525
qed
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   526
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   527
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   528
446
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   529
lemma rdistinct_equality1:
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   530
  shows "a \<notin> ss \<Longrightarrow> rdistinct (a  # rs) ss = a # rdistinct rs (insert a ss) "
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   531
  by auto
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   532
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   533
446
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   534
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   535
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   536
lemma larger_acc_smaller_distinct_res:
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   537
  shows " (sum_list (map rsize (rdistinct rs ss))) \<ge> (sum_list (map rsize (rdistinct rs (insert a ss))))"
447
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   538
  apply(subgoal_tac "ss \<subseteq> (insert a ss)")
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   539
   apply(rule larger_acc_smaller_distinct_res0)
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   540
   apply simp
7fb1e3dd5ae6 closedformbounds
Chengsong
parents: 446
diff changeset
   541
  by (simp add: subset_insertI)
446
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   542
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   543
lemma size_list_triangle1:
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   544
  shows  "sum_list (map rsize (a # (rdistinct as ss))) \<ge> rsize a + sum_list (map rsize (rdistinct as (insert a ss)))"
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   545
  by (simp add: larger_acc_smaller_distinct_res)
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   546
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   547
445
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   548
lemma triangle_inequality_distinct:
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   549
  shows "sum_list (map rsize (rdistinct (a # rs) ss)) \<le> rsize a + (sum_list (map rsize (rdistinct rs ss)))"
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   550
  apply(case_tac "a \<in> ss")
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   551
   apply simp
446
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   552
  apply(subst rdistinct_equality1)
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   553
   apply simp
0a94fd47b792 before repair
Chengsong
parents: 445
diff changeset
   554
  using size_list_triangle1 by auto
445
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   555
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   556
lemma same_regex_property_after_map:
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   557
  shows "\<forall>s. P (f r2 s) \<Longrightarrow> \<forall>r \<in> set  (map (f r2) Ss). P r"
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   558
  by auto
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   559
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   560
lemma same_property_after_distinct:
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   561
  shows " \<forall>r \<in> set  (map (f r2) Ss). P r \<Longrightarrow> \<forall>r \<in> set (rdistinct (map (f r2) Ss) xset). P r"
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   562
  apply(induct Ss arbitrary: xset)
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   563
   apply simp
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   564
  by auto
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   565
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   566
lemma same_regex_property_after_distinct:
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   567
  shows "\<forall>s. P (f r2 s) \<Longrightarrow> \<forall>r \<in> set (rdistinct (map (f r2) Ss) xset). P r"
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   568
  apply(rule same_property_after_distinct)
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   569
  apply(rule same_regex_property_after_map)
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   570
  by simp
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   571
449
Chengsong
parents: 448
diff changeset
   572
Chengsong
parents: 448
diff changeset
   573
Chengsong
parents: 448
diff changeset
   574
lemma Sum_cons:
Chengsong
parents: 448
diff changeset
   575
  shows "distinct (a # as) \<Longrightarrow> \<Sum> (set ((a::nat) # as)) =  a + \<Sum> (set  as)"
Chengsong
parents: 448
diff changeset
   576
  by simp
Chengsong
parents: 448
diff changeset
   577
Chengsong
parents: 448
diff changeset
   578
Chengsong
parents: 448
diff changeset
   579
lemma distinct_list_sizeNregex_bounded:
Chengsong
parents: 448
diff changeset
   580
  shows "distinct rs \<and> (\<forall> r \<in> (set rs). rsize r \<le> N) \<Longrightarrow> sum_list (map rsize rs) \<le> N * length rs"
Chengsong
parents: 448
diff changeset
   581
  apply(induct rs)
Chengsong
parents: 448
diff changeset
   582
   apply simp
Chengsong
parents: 448
diff changeset
   583
  by simp
Chengsong
parents: 448
diff changeset
   584
Chengsong
parents: 448
diff changeset
   585
Chengsong
parents: 448
diff changeset
   586
lemma distinct_list_size_len_bounded:
Chengsong
parents: 448
diff changeset
   587
  shows "distinct rs \<and> (\<forall>r \<in> set rs. rsize r \<le> N) \<and> length rs \<le> lrs \<Longrightarrow> sum_list (map rsize rs) \<le> lrs * N "
Chengsong
parents: 448
diff changeset
   588
  by (metis distinct_list_sizeNregex_bounded dual_order.trans mult.commute mult_le_mono1)
Chengsong
parents: 448
diff changeset
   589
Chengsong
parents: 448
diff changeset
   590
Chengsong
parents: 448
diff changeset
   591
Chengsong
parents: 448
diff changeset
   592
lemma rdistinct_same_set:
Chengsong
parents: 448
diff changeset
   593
  shows "(r \<in> set rs) =  (r \<in> set (rdistinct rs {}))"
Chengsong
parents: 448
diff changeset
   594
  apply(induct rs)
Chengsong
parents: 448
diff changeset
   595
   apply simp
Chengsong
parents: 448
diff changeset
   596
  apply(case_tac "a \<in> set rs")
Chengsong
parents: 448
diff changeset
   597
  apply(case_tac "r = a")
Chengsong
parents: 448
diff changeset
   598
    apply (simp)
Chengsong
parents: 448
diff changeset
   599
  apply (simp add: not_mentioned_elem_distinct)
Chengsong
parents: 448
diff changeset
   600
  using not_mentioned_elem_distinct by fastforce
Chengsong
parents: 448
diff changeset
   601
Chengsong
parents: 448
diff changeset
   602
Chengsong
parents: 448
diff changeset
   603
Chengsong
parents: 448
diff changeset
   604
lemma distinct_list_rexp_up_to_certain_size_bouded_by_set_enumerating_up_to_that_size:
Chengsong
parents: 448
diff changeset
   605
  shows "\<forall>r\<in> set rs. (rsize r ) \<le> N \<Longrightarrow> sum_list (map rsize (rdistinct rs {})) \<le>
Chengsong
parents: 448
diff changeset
   606
         (card (sizeNregex N))* N"
Chengsong
parents: 448
diff changeset
   607
  apply(subgoal_tac "distinct (rdistinct rs {})")
Chengsong
parents: 448
diff changeset
   608
  prefer 2
Chengsong
parents: 448
diff changeset
   609
  using rdistinct_does_the_job apply blast
Chengsong
parents: 448
diff changeset
   610
  apply(subgoal_tac "length (rdistinct rs {}) \<le> card (sizeNregex N)")
Chengsong
parents: 448
diff changeset
   611
  apply(rule distinct_list_size_len_bounded)
Chengsong
parents: 448
diff changeset
   612
   apply(rule conjI)+
Chengsong
parents: 448
diff changeset
   613
    apply simp
Chengsong
parents: 448
diff changeset
   614
   apply(rule conjI)
Chengsong
parents: 448
diff changeset
   615
  apply (meson rdistinct_same_set)
Chengsong
parents: 448
diff changeset
   616
   apply blast
Chengsong
parents: 448
diff changeset
   617
  apply(subgoal_tac "\<forall>r \<in> set (rdistinct rs {}). rsize r \<le> N")
Chengsong
parents: 448
diff changeset
   618
  prefer 2
Chengsong
parents: 448
diff changeset
   619
   apply (meson rdistinct_same_set)
Chengsong
parents: 448
diff changeset
   620
  apply(subgoal_tac "length (rdistinct rs {}) = card (set (rdistinct rs {}))")
Chengsong
parents: 448
diff changeset
   621
  prefer 2
Chengsong
parents: 448
diff changeset
   622
  using set_related_list apply blast
Chengsong
parents: 448
diff changeset
   623
  apply(simp only:)
Chengsong
parents: 448
diff changeset
   624
  by (metis card_mono finite_size_n mem_Collect_eq sizeNregex_def subset_code(1))
Chengsong
parents: 448
diff changeset
   625
Chengsong
parents: 448
diff changeset
   626
Chengsong
parents: 448
diff changeset
   627
Chengsong
parents: 448
diff changeset
   628
Chengsong
parents: 448
diff changeset
   629
Chengsong
parents: 448
diff changeset
   630
Chengsong
parents: 448
diff changeset
   631
lemma star_closed_form_bounded_by_rdistinct_list_estimate:
Chengsong
parents: 448
diff changeset
   632
  shows "rsize (rsimp ( RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) )
Chengsong
parents: 448
diff changeset
   633
         (star_updates s r0 [[c]]) ) ))) \<le>
Chengsong
parents: 448
diff changeset
   634
        Suc (sum_list (map rsize (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) )
Chengsong
parents: 448
diff changeset
   635
         (star_updates s r0 [[c]]) ) {})  ) )"
Chengsong
parents: 448
diff changeset
   636
  by (metis alts_simp_control )
Chengsong
parents: 448
diff changeset
   637
Chengsong
parents: 448
diff changeset
   638
Chengsong
parents: 448
diff changeset
   639
Chengsong
parents: 448
diff changeset
   640
Chengsong
parents: 448
diff changeset
   641
lemma star_lambda_form:
Chengsong
parents: 448
diff changeset
   642
  shows "\<forall> r \<in> set (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) ls). 
Chengsong
parents: 448
diff changeset
   643
        \<exists>s2. r = RSEQ (rders_simp r0 s2) (RSTAR r0) "
Chengsong
parents: 448
diff changeset
   644
  by (meson ex_map_conv)
Chengsong
parents: 448
diff changeset
   645
Chengsong
parents: 448
diff changeset
   646
Chengsong
parents: 448
diff changeset
   647
lemma star_lambda_ders:
Chengsong
parents: 448
diff changeset
   648
  shows " \<forall>s. rsize (rders_simp r0 s) \<le> N \<Longrightarrow>
Chengsong
parents: 448
diff changeset
   649
    \<forall>r\<in>set (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) (star_updates s r0 [[c]])).
Chengsong
parents: 448
diff changeset
   650
       rsize r \<le> Suc (N + rsize (RSTAR r0))"
Chengsong
parents: 448
diff changeset
   651
  apply(insert star_lambda_form)
Chengsong
parents: 448
diff changeset
   652
  apply(simp)
Chengsong
parents: 448
diff changeset
   653
  done
Chengsong
parents: 448
diff changeset
   654
Chengsong
parents: 448
diff changeset
   655
Chengsong
parents: 448
diff changeset
   656
Chengsong
parents: 448
diff changeset
   657
Chengsong
parents: 448
diff changeset
   658
lemma star_control_bounded:
Chengsong
parents: 448
diff changeset
   659
  shows "\<forall>s. rsize (rders_simp r0 s) \<le> N \<Longrightarrow>        
Chengsong
parents: 448
diff changeset
   660
      (sum_list (map rsize (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0))
Chengsong
parents: 448
diff changeset
   661
         (star_updates s r0 [[c]]) ) {})  ) ) \<le> 
Chengsong
parents: 448
diff changeset
   662
(card (sizeNregex (Suc (N + rsize (RSTAR r0))))) * (Suc (N + rsize (RSTAR r0)))
Chengsong
parents: 448
diff changeset
   663
"
Chengsong
parents: 448
diff changeset
   664
  apply(subgoal_tac "\<forall>r \<in> set (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0))
Chengsong
parents: 448
diff changeset
   665
         (star_updates s r0 [[c]]) ). (rsize r ) \<le> Suc (N + rsize (RSTAR r0))")
Chengsong
parents: 448
diff changeset
   666
   prefer 2
Chengsong
parents: 448
diff changeset
   667
  using star_lambda_ders apply blast
Chengsong
parents: 448
diff changeset
   668
  using distinct_list_rexp_up_to_certain_size_bouded_by_set_enumerating_up_to_that_size by blast
Chengsong
parents: 448
diff changeset
   669
Chengsong
parents: 448
diff changeset
   670
Chengsong
parents: 448
diff changeset
   671
lemma star_control_variant:
Chengsong
parents: 448
diff changeset
   672
  assumes "\<forall>s. rsize (rders_simp r0 s) \<le> N"
Chengsong
parents: 448
diff changeset
   673
  shows"Suc 
Chengsong
parents: 448
diff changeset
   674
      (sum_list (map rsize (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) 
Chengsong
parents: 448
diff changeset
   675
          (star_updates list r0 [[a]])) {}))) 
Chengsong
parents: 448
diff changeset
   676
\<le>  (Suc (card (sizeNregex (Suc (N + rsize (RSTAR r0))))) * Suc (N + rsize (RSTAR r0))) "
Chengsong
parents: 448
diff changeset
   677
  apply(subgoal_tac    "(sum_list (map rsize (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) 
Chengsong
parents: 448
diff changeset
   678
          (star_updates list r0 [[a]])) {}))) 
Chengsong
parents: 448
diff changeset
   679
\<le>  ( (card (sizeNregex (Suc (N + rsize (RSTAR r0))))) * Suc (N + rsize (RSTAR r0))) ")
Chengsong
parents: 448
diff changeset
   680
  prefer 2
Chengsong
parents: 448
diff changeset
   681
  using assms star_control_bounded apply presburger
Chengsong
parents: 448
diff changeset
   682
  by simp
Chengsong
parents: 448
diff changeset
   683
Chengsong
parents: 448
diff changeset
   684
Chengsong
parents: 448
diff changeset
   685
Chengsong
parents: 448
diff changeset
   686
lemma star_closed_form_bounded:
Chengsong
parents: 448
diff changeset
   687
  shows "\<forall>s. rsize (rders_simp r0 s) \<le> N \<Longrightarrow>
Chengsong
parents: 448
diff changeset
   688
              rsize (rders_simp (RSTAR r0) s) \<le> 
Chengsong
parents: 448
diff changeset
   689
max (   (Suc (card (sizeNregex (Suc (N + rsize (RSTAR r0))))) * (Suc (N + rsize (RSTAR r0)))))   (rsize (RSTAR r0))"
Chengsong
parents: 448
diff changeset
   690
  apply(case_tac s)
Chengsong
parents: 448
diff changeset
   691
  apply simp
Chengsong
parents: 448
diff changeset
   692
  apply(subgoal_tac " rsize (rders_simp (RSTAR r0) (a # list)) = 
Chengsong
parents: 448
diff changeset
   693
rsize (rsimp ( RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) (star_updates list r0 [[a]]) ) )))") 
Chengsong
parents: 448
diff changeset
   694
   prefer 2
Chengsong
parents: 448
diff changeset
   695
  using star_closed_form apply presburger
Chengsong
parents: 448
diff changeset
   696
  apply(subgoal_tac "rsize (rsimp (
Chengsong
parents: 448
diff changeset
   697
 RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) (star_updates list    r0 [[a]]) ) ))) 
Chengsong
parents: 448
diff changeset
   698
\<le>         Suc (sum_list (map rsize (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) )
Chengsong
parents: 448
diff changeset
   699
         (star_updates list r0 [[a]]) ) {})  ) )")
Chengsong
parents: 448
diff changeset
   700
  prefer 2
Chengsong
parents: 448
diff changeset
   701
  using star_closed_form_bounded_by_rdistinct_list_estimate apply presburger
Chengsong
parents: 448
diff changeset
   702
  apply(subgoal_tac "Suc (sum_list 
Chengsong
parents: 448
diff changeset
   703
                 (map rsize
Chengsong
parents: 448
diff changeset
   704
                   (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) (star_updates list r0 [[a]])) {}))) 
Chengsong
parents: 448
diff changeset
   705
\<le>  (Suc (card (sizeNregex (Suc (N + rsize (RSTAR r0))))) * Suc (N + rsize (RSTAR r0)))  ")
Chengsong
parents: 448
diff changeset
   706
  apply auto[1]
Chengsong
parents: 448
diff changeset
   707
  using star_control_variant by blast
Chengsong
parents: 448
diff changeset
   708
Chengsong
parents: 448
diff changeset
   709
Chengsong
parents: 448
diff changeset
   710
lemma seq_list_estimate_control: shows 
Chengsong
parents: 448
diff changeset
   711
" rsize (rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1))))
Chengsong
parents: 448
diff changeset
   712
           \<le> Suc (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {})))"
Chengsong
parents: 448
diff changeset
   713
  by(metis alts_simp_control)
Chengsong
parents: 448
diff changeset
   714
445
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   715
lemma map_ders_is_list_of_ders:
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   716
  shows  "\<forall>s. rsize (rders_simp r2 s) \<le> N2 \<Longrightarrow>
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   717
\<forall>r \<in> set (rdistinct (map (rders_simp r2) Ss) {}). rsize r \<le> N2"
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   718
  apply(rule same_regex_property_after_distinct)
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   719
  by simp
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   720
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   721
lemma seq_estimate_bounded: 
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   722
  assumes "\<forall>s. rsize (rders_simp r1 s) \<le> N1" and "\<forall>s. rsize (rders_simp r2 s) \<le> N2"
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   723
  shows
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   724
"Suc (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {}))) \<le>
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   725
 Suc (Suc (N1 + (rsize r2)) + (N2 * card (sizeNregex N2)))"
445
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   726
  apply(subgoal_tac " (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {}))) \<le>
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   727
  (Suc (N1 + (rsize r2)) + (N2 * card (sizeNregex N2)))")
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   728
   apply force
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   729
  apply(subgoal_tac " (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {}))) \<le>
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   730
                      (rsize (RSEQ (rders_simp r1 s) r2)) + (sum_list (map rsize (rdistinct (map (rders_simp r2) (vsuf s r1)) {})) )")
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   731
  prefer 2
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   732
  using triangle_inequality_distinct apply blast
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   733
  apply(subgoal_tac " sum_list (map rsize (rdistinct (map (rders_simp r2) (vsuf s r1)) {})) \<le> N2 * card (sizeNregex N2) ")
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   734
   apply(subgoal_tac "rsize (RSEQ (rders_simp r1 s) r2) \<le> Suc (N1 + rsize r2)")
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   735
    apply linarith
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   736
   apply (simp add: assms(1))
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   737
  apply(subgoal_tac "\<forall>r \<in> set (rdistinct (map (rders_simp r2) (vsuf s r1)) {}). rsize r \<le> N2")
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   738
  apply (metis (no_types, opaque_lifting) assms(2) distinct_list_rexp_up_to_certain_size_bouded_by_set_enumerating_up_to_that_size ex_map_conv mult.commute)
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
   739
  using assms(2) map_ders_is_list_of_ders by blast
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   740
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   741
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   742
lemma seq_closed_form_bounded: shows
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   743
"\<lbrakk>\<forall>s. rsize (rders_simp r1 s) \<le> N1 ; \<forall>s. rsize (rders_simp r2 s) \<le> N2\<rbrakk> \<Longrightarrow>
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   744
rsize (rders_simp (RSEQ r1 r2) s) \<le> 
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   745
max (Suc (Suc (N1 + (rsize r2)) + (N2 * card (sizeNregex N2)))) (rsize (RSEQ r1 r2)) "
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   746
  apply(case_tac s)
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   747
  apply simp
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   748
  apply(subgoal_tac " (rders_simp (RSEQ r1 r2) s) = 
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   749
rsimp (RALTS ((RSEQ (rders_simp r1 s) r2) # (map (rders_simp r2) (vsuf s r1))))")
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   750
  prefer 2
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   751
  using seq_closed_form_variant apply blast
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   752
  apply(subgoal_tac "rsize (rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1))))
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   753
                    \<le>
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   754
Suc (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {})))")
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   755
  apply(subgoal_tac "Suc (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {})))
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   756
\<le> Suc (Suc (N1 + (rsize r2)) + (N2 * card (sizeNregex N2)))")
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   757
  prefer 2
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   758
  using seq_estimate_bounded apply blast
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   759
   apply(subgoal_tac "rsize (rders_simp (RSEQ r1 r2) s) \<le> Suc (Suc (N1 + rsize r2) + N2 * card (sizeNregex N2))")
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   760
  using le_max_iff_disj apply blast
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   761
   apply auto[1]
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   762
  using seq_list_estimate_control by presburger
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   763
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   764
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   765
lemma rders_simp_bounded: shows
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   766
"\<exists>N. \<forall>s. rsize (rders_simp r s) \<le> N"
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   767
  apply(induct r)
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   768
       apply(rule_tac x = "Suc 0 " in exI)
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   769
  using three_easy_cases0 apply force
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   770
  using three_easy_cases1 apply blast
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   771
  using three_easy_casesC apply blast
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   772
  using seq_closed_form_bounded apply blast
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   773
  apply (metis alts_closed_form_bounded size_list_estimation')
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   774
  using star_closed_form_bounded by blast
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   775
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   776
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   777
451
7a016eeb118d finiteness
Chengsong
parents: 450
diff changeset
   778
unused_thms
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   779
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   780
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   781
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   782
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   783
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   784
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   785
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   786
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   787
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   788
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   789
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   790
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   791
(*Obsolete materials*)
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   792
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   793
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   794
a7e98deebb5c restructured sizebound proof
Chengsong
parents:
diff changeset
   795
end