isahipster: e81927c04db1762e324f7e313a1e95c29d05de6e

     1: theory prop_14
     2: imports Main
     3:         "$HIPSTER_HOME/IsaHipster"
     4: begin
     5:   datatype 'a list = Nil2 | Cons2 "'a" "'a list"
     6:   datatype Nat = Z | S "Nat"
     7:   fun le :: "Nat => Nat => bool" where
     8:   "le (Z) y = True"
     9:   | "le (S z) (Z) = False"
    10:   | "le (S z) (S x2) = le z x2"
    11:   fun sorted :: "Nat list => bool" where
    12:   "sorted (Nil2) = True"
    13:   | "sorted (Cons2 y (Nil2)) = True"
    14:   | "sorted (Cons2 y (Cons2 y2 xs)) =
    15:        (if le y y2 then sorted (Cons2 y2 xs) else False)"
    16:   fun insert2 :: "Nat => Nat list => Nat list" where
    17:   "insert2 x (Nil2) = Cons2 x (Nil2)"
    18:   | "insert2 x (Cons2 z xs) =
    19:        (if le x z then Cons2 x (Cons2 z xs) else Cons2 z (insert2 x xs))"
    20:   fun isort :: "Nat list => Nat list" where
    21:   "isort (Nil2) = Nil2"
    22:   | "isort (Cons2 y xs) = insert2 y (isort xs)"
    23:   (*hipster le sorted insert2 isort *)
    24: 
    25: hipster le
    26: lemma lemma_a [thy_expl]: "le x2 x2 = True"
    27: by (hipster_induct_schemes le.simps)
    28: 
    29: lemma lemma_aa [thy_expl]: "le x2 (S x2) = True"
    30: by (hipster_induct_schemes le.simps)
    31: 
    32: lemma lemma_ab [thy_expl]: "le (S x2) x2 = False"
    33: by (hipster_induct_schemes le.simps)
    34: 
    35: (*hipster_cond le*)
    36: lemma lemma_ac [thy_expl]: "le x2 y2 \<Longrightarrow> le x2 (S y2) = True"
    37: by (hipster_induct_schemes le.simps)
    38: 
    39: lemma lemma_ad [thy_expl]: "le y2 x2 \<Longrightarrow> le (S x2) y2 = False"
    40: by (hipster_induct_schemes le.simps)
    41: 
    42: lemma lemma_ae [thy_expl]: "le y x \<and> le x y \<Longrightarrow> x = y"
    43: by (hipster_induct_schemes le.simps Nat.exhaust)
    44: 
    45: lemma le_trans [thy_expl]: "le z y \<and> le x z \<Longrightarrow> le x y = True"
    46: by (hipster_induct_schemes le.simps Nat.exhaust)
    47: 
    48: 
    49: (*hipster_cond sorted isort insert2 le*)
    50: lemma lemma_af [thy_expl]: "insert2 Z (insert2 x20 y20) = insert2 x20 (insert2 Z y20)"
    51: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps le.simps)
    52: 
    53: lemma lemma_ag [thy_expl]: "sorted (insert2 Z x4) = sorted x4"
    54: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps le.simps)
    55: 
    56: lemma lemma_ah [thy_expl]: "insert2 Z (isort x3) = isort (insert2 Z x3)"
    57: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps le.simps)
    58: 
    59: (* include negation: discovered if declared separately *)
    60: lemma nole [thy_expl]: "\<not> le r s \<Longrightarrow> le s r"
    61: by (hipster_induct_schemes le.simps Nat.exhaust)
    62: 
    63: (*hipster_cond sorted isort insert2 le*)
    64: lemma lemma_ai [thy_expl]: "insert2 x32 (insert2 y32 z32) = insert2 y32 (insert2 x32 z32)"
    65: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps le.simps)
    66: 
    67: lemma lemma_aj [thy_expl]: "isort (insert2 x19 y19) = insert2 x19 (isort y19)"
    68: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps le.simps)
    69: 
    70: lemma lemma_ak [thy_expl]: "isort (isort x19) = isort x19"
    71: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps le.simps)
    72: 
    73: setup {* Hip_Tac_Ops.set_full_types @{context} true *}
    74: setup {* Hip_Tac_Ops.set_metis_to @{context} 700 *}
    75: 
    76: lemma lemma_al [thy_expl]: "sorted x \<Longrightarrow> isort x = x"
    77: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps le.simps)
    78: 
    79: setup {* Hip_Tac_Ops.set_metis_to @{context} 1000 *}
    80: 
    81: lemma lemma_am [thy_expl]: "sorted y \<Longrightarrow> sorted (insert2 x y) = True"
    82: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps nole)
    83: 
    84: lemma lemma_an [thy_expl]: "sorted (isort x) = True"
    85: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps)
    86: 
    87: lemma lemma_ao [thy_expl]: "sorted y \<Longrightarrow> isort (insert2 x y) = insert2 x y"
    88: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps)
    89: (*
    90: lemma lemma_ap [thy_expl]: "sorted (insert2 x y) = sorted y"
    91: by (hipster_induct_schemes sorted.simps isort.simps insert2.simps le.simps Nat.exhaust)*)
    92: 
    93: 
    94:   theorem x0 :
    95:     "sorted (isort x)"
    96:     by (tactic {* Subgoal.FOCUS_PARAMS (K (Tactic_Data.hard_tac @{context})) @{context} 1 *})
    97: end

Generated by git2html.