src/Tools/isac/Knowledge/Biegelinie.thy
author Walther Neuper <walther.neuper@jku.at>
Wed, 15 Apr 2020 11:37:43 +0200
changeset 59878 3163e63a5111
parent 59871 82428ca0d23e
child 59898 68883c046963
permissions -rw-r--r--
cleanup
neuper@37906
     1
(* chapter 'Biegelinie' from the textbook: 
neuper@37906
     2
   Timischl, Kaiser. Ingenieur-Mathematik 3. Wien 1999. p.268-271.
neuper@37999
     3
   author: Walther Neuper 050826,
neuper@37906
     4
   (c) due to copyright terms
neuper@37906
     5
*)
neuper@37906
     6
wneuper@59424
     7
theory Biegelinie imports Integrate Equation EqSystem Base_Tools begin
neuper@37906
     8
neuper@37906
     9
consts
neuper@37906
    10
neuper@42385
    11
  qq    :: "real => real"             (* Streckenlast               *)
neuper@37999
    12
  Q     :: "real => real"             (* Querkraft                  *)
neuper@37999
    13
  Q'    :: "real => real"             (* Ableitung der Querkraft    *)
neuper@37999
    14
  M'_b  :: "real => real" ("M'_b")    (* Biegemoment                *)
neuper@37999
    15
  M'_b' :: "real => real" ("M'_b'")   (* Ableitung des Biegemoments *)
neuper@37999
    16
  y''   :: "real => real"             (* 2.Ableitung der Biegeline  *)
neuper@37999
    17
  y'    :: "real => real"             (* Neigung der Biegeline      *)
neuper@37999
    18
(*y     :: "real => real"             (* Biegeline                  *)*)
neuper@37999
    19
  EI    :: "real"                     (* Biegesteifigkeit           *)
neuper@37906
    20
neuper@37906
    21
  (*new Descriptions in the related problems*)
neuper@37999
    22
  Traegerlaenge            :: "real => una"
neuper@37999
    23
  Streckenlast             :: "real => una"
neuper@37999
    24
  BiegemomentVerlauf       :: "bool => una"
neuper@37999
    25
  Biegelinie               :: "(real => real) => una"
wneuper@59540
    26
  AbleitungBiegelinie      :: "(real => real) => una"
neuper@37999
    27
  Randbedingungen          :: "bool list => una"
neuper@37999
    28
  RandbedingungenBiegung   :: "bool list => una"
neuper@37999
    29
  RandbedingungenNeigung   :: "bool list => una"
neuper@37999
    30
  RandbedingungenMoment    :: "bool list => una"
neuper@37999
    31
  RandbedingungenQuerkraft :: "bool list => una"
neuper@37999
    32
  FunktionsVariable        :: "real => una"
neuper@37999
    33
  Funktionen               :: "bool list => una"
neuper@37999
    34
  Gleichungen              :: "bool list => una"
wneuper@59539
    35
  GleichungsVariablen      :: "real list => una"
neuper@37906
    36
neuper@52148
    37
axiomatization where
neuper@37906
    38
neuper@52148
    39
  Querkraft_Belastung:   "Q' x = -qq x" and
neuper@52148
    40
  Belastung_Querkraft:   "-qq x = Q' x" and
neuper@37906
    41
neuper@52148
    42
  Moment_Querkraft:      "M_b' x = Q x" and
neuper@52148
    43
  Querkraft_Moment:      "Q x = M_b' x" and
neuper@37906
    44
neuper@52148
    45
  Neigung_Moment:        "y'' x = -M_b x/ EI" and
neuper@52148
    46
  Moment_Neigung:        "M_b x = -EI * y'' x" and
neuper@37906
    47
neuper@37906
    48
  (*according to rls 'simplify_Integral': .. = 1/a * .. instead .. = ../ a*)
neuper@37983
    49
  make_fun_explicit:     "Not (a =!= 0) ==> (a * (f x) = b) = (f x = 1/a * b)"
neuper@37906
    50
wneuper@59472
    51
setup \<open>
neuper@55425
    52
KEStore_Elems.add_thes
neuper@55425
    53
  [make_thy @{theory}
neuper@55425
    54
    ["Walther Neuper 2005 supported by a grant from NMI Austria"],
neuper@55425
    55
  make_isa @{theory} ("IsacKnowledge", "Theorems")
neuper@55425
    56
    ["Walther Neuper 2005 supported by a grant from NMI Austria"],
neuper@55489
    57
  make_thm @{theory}  "IsacKnowledge" ("Belastung_Querkraft", @{thm Belastung_Querkraft})
neuper@55425
    58
	  ["Walther Neuper 2005 supported by a grant from NMI Austria"],
neuper@55489
    59
  make_thm @{theory}  "IsacKnowledge" ("Moment_Neigung", @{thm Moment_Neigung})
neuper@55425
    60
	  ["Walther Neuper 2005 supported by a grant from NMI Austria"],
neuper@55489
    61
  make_thm @{theory}  "IsacKnowledge" ("Moment_Querkraft", @{thm Moment_Querkraft})
neuper@55425
    62
	  ["Walther Neuper 2005 supported by a grant from NMI Austria"],
neuper@55489
    63
  make_thm @{theory}  "IsacKnowledge" ("Neigung_Moment", @{thm Neigung_Moment})
neuper@55425
    64
	  ["Walther Neuper 2005 supported by a grant from NMI Austria"],
neuper@55489
    65
  make_thm @{theory}  "IsacKnowledge" ("Querkraft_Belastung", @{thm Querkraft_Belastung})
neuper@55425
    66
	  ["Walther Neuper 2005 supported by a grant from NMI Austria"],
neuper@55489
    67
  make_thm @{theory}  "IsacKnowledge" ("Querkraft_Moment", @{thm Querkraft_Moment})
neuper@55425
    68
	  ["Walther Neuper 2005 supported by a grant from NMI Austria"],
neuper@55489
    69
  make_thm @{theory}  "IsacKnowledge" ("make_fun_explicit", @{thm make_fun_explicit})
neuper@55425
    70
	  ["Walther Neuper 2005 supported by a grant from NMI Austria"]]
wneuper@59472
    71
\<close>
neuper@37954
    72
neuper@37954
    73
(** problems **)
wneuper@59472
    74
setup \<open>KEStore_Elems.add_pbts
wneuper@59406
    75
  [(Specify.prep_pbt @{theory} "pbl_bieg" [] Celem.e_pblID
s1210629013@55339
    76
      (["Biegelinien"],
s1210629013@55339
    77
        [("#Given" ,["Traegerlaenge l_l", "Streckenlast q_q"]),
s1210629013@55339
    78
          (*("#Where",["0 < l_l"]), ...wait for &lt; and handling Arbfix*)
s1210629013@55339
    79
          ("#Find"  ,["Biegelinie b_b"]),
s1210629013@55339
    80
          ("#Relate",["Randbedingungen r_b"])],
walther@59852
    81
        Rule_Set.append_rules "empty" Rule_Set.empty [], NONE, [["IntegrierenUndKonstanteBestimmen2"]])),
wneuper@59406
    82
    (Specify.prep_pbt @{theory} "pbl_bieg_mom" [] Celem.e_pblID
s1210629013@55339
    83
      (["MomentBestimmte","Biegelinien"],
s1210629013@55339
    84
        [("#Given" ,["Traegerlaenge l_l", "Streckenlast q_q"]),
s1210629013@55339
    85
          (*("#Where",["0 < l_l"]), ...wait for &lt; and handling Arbfix*)
s1210629013@55339
    86
          ("#Find"  ,["Biegelinie b_b"]),
s1210629013@55339
    87
          ("#Relate",["RandbedingungenBiegung r_b","RandbedingungenMoment r_m"])
s1210629013@55339
    88
        ],
walther@59852
    89
        Rule_Set.append_rules "empty" Rule_Set.empty [], NONE, [["IntegrierenUndKonstanteBestimmen"]])),
wneuper@59406
    90
    (Specify.prep_pbt @{theory} "pbl_bieg_momg" [] Celem.e_pblID
walther@59852
    91
      (["MomentGegebene","Biegelinien"], [], Rule_Set.append_rules "empty" Rule_Set.empty [], NONE,
s1210629013@55339
    92
        [["IntegrierenUndKonstanteBestimmen","2xIntegrieren"]])),
wneuper@59406
    93
    (Specify.prep_pbt @{theory} "pbl_bieg_einf" [] Celem.e_pblID
walther@59852
    94
      (["einfache","Biegelinien"], [], Rule_Set.append_rules "empty" Rule_Set.empty [], NONE,
s1210629013@55339
    95
        [["IntegrierenUndKonstanteBestimmen","4x4System"]])),
wneuper@59406
    96
    (Specify.prep_pbt @{theory} "pbl_bieg_momquer" [] Celem.e_pblID
walther@59852
    97
      (["QuerkraftUndMomentBestimmte","Biegelinien"], [], Rule_Set.append_rules "empty" Rule_Set.empty [], NONE,
s1210629013@55339
    98
        [["IntegrierenUndKonstanteBestimmen","1xIntegrieren"]])),
wneuper@59406
    99
    (Specify.prep_pbt @{theory} "pbl_bieg_vonq" [] Celem.e_pblID
s1210629013@55339
   100
      (["vonBelastungZu","Biegelinien"],
s1210629013@55339
   101
          [("#Given" ,["Streckenlast q_q","FunktionsVariable v_v"]),
s1210629013@55339
   102
            ("#Find"  ,["Funktionen funs'''"])],
walther@59852
   103
        Rule_Set.append_rules "empty" Rule_Set.empty [], NONE, [["Biegelinien","ausBelastung"]])),
wneuper@59406
   104
    (Specify.prep_pbt @{theory} "pbl_bieg_randbed" [] Celem.e_pblID
s1210629013@55339
   105
      (["setzeRandbedingungen","Biegelinien"],
s1210629013@55339
   106
          [("#Given" ,["Funktionen fun_s","Randbedingungen r_b"]),
s1210629013@55339
   107
            ("#Find"  ,["Gleichungen equs'''"])],
walther@59852
   108
        Rule_Set.append_rules "empty" Rule_Set.empty [], NONE, [["Biegelinien","setzeRandbedingungenEin"]])),
wneuper@59406
   109
    (Specify.prep_pbt @{theory} "pbl_equ_fromfun" [] Celem.e_pblID
s1210629013@55339
   110
      (["makeFunctionTo","equation"],
s1210629013@55339
   111
          [("#Given" ,["functionEq fu_n","substitution su_b"]),
s1210629013@55339
   112
            ("#Find"  ,["equality equ'''"])],
walther@59852
   113
        Rule_Set.append_rules "empty" Rule_Set.empty [], NONE, [["Equation","fromFunction"]]))]\<close>
wneuper@59472
   114
ML \<open>
neuper@37954
   115
(** methods **)
neuper@37954
   116
walther@59851
   117
val srls = Rule_Def.Repeat {id="srls_IntegrierenUnd..", 
neuper@37954
   118
		preconds = [], 
neuper@37954
   119
		rew_ord = ("termlessI",termlessI), 
walther@59852
   120
		erls = Rule_Set.append_rules "erls_in_srls_IntegrierenUnd.." Rule_Set.empty
neuper@37999
   121
				  [(*for asm in NTH_CONS ...*)
walther@59878
   122
				   Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
neuper@37999
   123
				   (*2nd NTH_CONS pushes n+-1 into asms*)
walther@59878
   124
				   Rule.Eval("Groups.plus_class.plus", (**)eval_binop "#add_")
neuper@37954
   125
				   ], 
walther@59851
   126
		srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59871
   127
		rules = [Rule.Thm ("NTH_CONS",ThmC.numerals_to_Free @{thm NTH_CONS}),
walther@59878
   128
			 Rule.Eval("Groups.plus_class.plus", (**)eval_binop "#add_"),
walther@59871
   129
			 Rule.Thm ("NTH_NIL",ThmC.numerals_to_Free @{thm NTH_NIL}),
walther@59878
   130
			 Rule.Eval("Prog_Expr.lhs", Prog_Expr.eval_lhs"eval_lhs_"),
walther@59878
   131
			 Rule.Eval("Prog_Expr.rhs", Prog_Expr.eval_rhs"eval_rhs_"),
walther@59878
   132
			 Rule.Eval("Prog_Expr.argument'_in", Prog_Expr.eval_argument_in "Prog_Expr.argument'_in")
neuper@37954
   133
			 ],
walther@59878
   134
		scr = Rule.Empty_Prog};
neuper@37954
   135
    
neuper@37954
   136
val srls2 = 
walther@59851
   137
    Rule_Def.Repeat {id="srls_IntegrierenUnd..", 
neuper@37954
   138
	 preconds = [], 
neuper@37954
   139
	 rew_ord = ("termlessI",termlessI), 
walther@59852
   140
	 erls = Rule_Set.append_rules "erls_in_srls_IntegrierenUnd.." Rule_Set.empty
neuper@37999
   141
			   [(*for asm in NTH_CONS ...*)
walther@59878
   142
			    Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
neuper@37999
   143
			    (*2nd NTH_CONS pushes n+-1 into asms*)
walther@59878
   144
			    Rule.Eval("Groups.plus_class.plus", (**)eval_binop "#add_")
neuper@37954
   145
			    ], 
walther@59851
   146
	 srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59871
   147
	 rules = [Rule.Thm ("NTH_CONS",ThmC.numerals_to_Free @{thm NTH_CONS}),
walther@59878
   148
		  Rule.Eval("Groups.plus_class.plus", (**)eval_binop "#add_"),
walther@59871
   149
		  Rule.Thm ("NTH_NIL", ThmC.numerals_to_Free @{thm NTH_NIL}),
walther@59878
   150
		  Rule.Eval("Prog_Expr.lhs", Prog_Expr.eval_lhs "eval_lhs_"),
walther@59878
   151
		  Rule.Eval("Prog_Expr.filter'_sameFunId", Prog_Expr.eval_filter_sameFunId "Prog_Expr.filter'_sameFunId"),
neuper@37954
   152
		  (*WN070514 just for smltest/../biegelinie.sml ...*)
walther@59878
   153
		  Rule.Eval("Prog_Expr.sameFunId", Prog_Expr.eval_sameFunId "Prog_Expr.sameFunId"),
walther@59871
   154
		  Rule.Thm ("filter_Cons", ThmC.numerals_to_Free @{thm filter_Cons}),
walther@59871
   155
		  Rule.Thm ("filter_Nil", ThmC.numerals_to_Free @{thm filter_Nil}),
walther@59871
   156
		  Rule.Thm ("if_True", ThmC.numerals_to_Free @{thm if_True}),
walther@59871
   157
		  Rule.Thm ("if_False", ThmC.numerals_to_Free @{thm if_False}),
walther@59871
   158
		  Rule.Thm ("hd_thm", ThmC.numerals_to_Free @{thm hd_thm})
neuper@37954
   159
		  ],
walther@59878
   160
	 scr = Rule.Empty_Prog};
wneuper@59472
   161
\<close>
neuper@37999
   162
wneuper@59429
   163
section \<open>Methods\<close>
wneuper@59506
   164
(* setup parent directory *)
wneuper@59472
   165
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   166
    [Specify.prep_met @{theory} "met_biege" [] Celem.e_metID 
wneuper@59506
   167
	    (["IntegrierenUndKonstanteBestimmen"], [],
walther@59851
   168
	    {rew_ord'="tless_true", rls'= Rule_Set.Empty, calc = [], srls = Rule_Set.Empty, prls = Rule_Set.Empty, crls =Rule_Set.Empty,
walther@59851
   169
          errpats = [], nrls = Rule_Set.Empty},
wneuper@59545
   170
      @{thm refl})]
wneuper@59472
   171
\<close>
wneuper@59429
   172
subsection \<open>Sub-problem "integrate and determine constants", nicely modularised\<close>
wneuper@59545
   173
wneuper@59549
   174
partial_function (tailrec) biegelinie ::
wneuper@59549
   175
  "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> bool list \<Rightarrow> real list \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> bool"
wneuper@59504
   176
  where
walther@59635
   177
"biegelinie l q v b s vs id_abl = (
walther@59635
   178
  let
wneuper@59504
   179
    funs = SubProblem (''Biegelinie'', [''vonBelastungZu'', ''Biegelinien''],
walther@59635
   180
      [''Biegelinien'', ''ausBelastung'']) [REAL q, REAL v, REAL_REAL b, REAL_REAL id_abl];
wneuper@59504
   181
    equs = SubProblem (''Biegelinie'', [''setzeRandbedingungen'', ''Biegelinien''],
walther@59635
   182
      [''Biegelinien'', ''setzeRandbedingungenEin'']) [BOOL_LIST funs, BOOL_LIST s];
walther@59635
   183
    cons = SubProblem (''Biegelinie'', [''LINEAR'', ''system''],
walther@59635
   184
      [''no_met'']) [BOOL_LIST equs, REAL_LIST vs];
wneuper@59504
   185
    B = Take (lastI funs);
walther@59637
   186
    B = ((Substitute cons) #> (Rewrite_Set_Inst [(''bdv'', v)] ''make_ratpoly_in'')) B
wneuper@59504
   187
  in B)"
wneuper@59472
   188
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   189
    [Specify.prep_met @{theory} "met_biege_2" [] Celem.e_metID
s1210629013@55373
   190
	    (["IntegrierenUndKonstanteBestimmen2"],
wneuper@59539
   191
	      [("#Given" ,["Traegerlaenge l", "Streckenlast q",
wneuper@59540
   192
              "FunktionsVariable v", "GleichungsVariablen vs", "AbleitungBiegelinie id_abl"]),
wneuper@59429
   193
		      (*("#Where",["0 < l"]), ...wait for &lt; and handling Arbfix*)
wneuper@59429
   194
		      ("#Find"  ,["Biegelinie b"]),
wneuper@59429
   195
		      ("#Relate",["Randbedingungen s"])],
s1210629013@55373
   196
	      {rew_ord'="tless_true", 
walther@59852
   197
	        rls' = Rule_Set.append_rules "erls_IntegrierenUndK.." Rule_Set.empty 
walther@59878
   198
				      [Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"),
walther@59871
   199
				        Rule.Thm ("not_true",ThmC.numerals_to_Free @{thm not_true}),
walther@59871
   200
				        Rule.Thm ("not_false",ThmC.numerals_to_Free @{thm not_false})], 
s1210629013@55373
   201
				  calc = [], 
walther@59852
   202
				  srls = Rule_Set.append_rules "erls_IntegrierenUndK.." Rule_Set.empty 
walther@59878
   203
				      [Rule.Eval("Prog_Expr.rhs", Prog_Expr.eval_rhs"eval_rhs_"),
walther@59878
   204
				        Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"),
walther@59871
   205
				        Rule.Thm ("last_thmI",ThmC.numerals_to_Free @{thm last_thmI}),
walther@59871
   206
				        Rule.Thm ("if_True",ThmC.numerals_to_Free @{thm if_True}),
walther@59871
   207
				        Rule.Thm ("if_False",ThmC.numerals_to_Free @{thm if_False})],
walther@59851
   208
				  prls = Rule_Set.Empty, crls = Atools_erls, errpats = [], nrls = Rule_Set.Empty},
wneuper@59551
   209
        @{thm biegelinie.simps})]
wneuper@59473
   210
\<close>
wneuper@59473
   211
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   212
    [Specify.prep_met @{theory} "met_biege_intconst_2" [] Celem.e_metID
s1210629013@55373
   213
	    (["IntegrierenUndKonstanteBestimmen","2xIntegrieren"], [],
walther@59852
   214
	      {rew_ord'="tless_true", rls'=Rule_Set.Empty, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty, crls = Atools_erls,
walther@59852
   215
          errpats = [], nrls = Rule_Set.empty},
wneuper@59545
   216
        @{thm refl}),
wneuper@59406
   217
    Specify.prep_met @{theory} "met_biege_intconst_4" [] Celem.e_metID
s1210629013@55373
   218
	    (["IntegrierenUndKonstanteBestimmen","4x4System"], [],
walther@59852
   219
	      {rew_ord'="tless_true", rls'=Rule_Set.Empty, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty, crls = Atools_erls,
walther@59852
   220
          errpats = [], nrls = Rule_Set.empty},
wneuper@59545
   221
        @{thm refl}),
wneuper@59406
   222
    Specify.prep_met @{theory} "met_biege_intconst_1" [] Celem.e_metID
s1210629013@55373
   223
	    (["IntegrierenUndKonstanteBestimmen","1xIntegrieren"], [],
walther@59852
   224
        {rew_ord'="tless_true", rls'=Rule_Set.Empty, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty, crls = Atools_erls,
walther@59852
   225
          errpats = [], nrls = Rule_Set.empty},
wneuper@59545
   226
        @{thm refl}),
wneuper@59406
   227
    Specify.prep_met @{theory} "met_biege2" [] Celem.e_metID
s1210629013@55373
   228
	    (["Biegelinien"], [],
walther@59852
   229
	      {rew_ord'="tless_true", rls'=Rule_Set.Empty, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty, crls = Atools_erls,
walther@59852
   230
          errpats = [], nrls = Rule_Set.empty},
wneuper@59545
   231
        @{thm refl})]
wneuper@59472
   232
\<close>
wneuper@59429
   233
subsection \<open>Compute the general bending line\<close>
wneuper@59429
   234
wneuper@59544
   235
partial_function (tailrec) belastung_zu_biegelinie :: "real \<Rightarrow> real \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> bool list"
wneuper@59504
   236
  where
walther@59635
   237
"belastung_zu_biegelinie q__q v_v id_fun id_abl = (
walther@59635
   238
  let
walther@59635
   239
    q___q = Take (qq v_v = q__q);
walther@59635
   240
    q___q = (
walther@59637
   241
      (Rewrite ''sym_neg_equal_iff_equal'') #>
walther@59635
   242
      (Rewrite ''Belastung_Querkraft'')) q___q;
walther@59635
   243
    Q__Q = SubProblem (''Biegelinie'', [''named'', ''integrate'', ''function''],
walther@59635
   244
      [''diff'', ''integration'', ''named'']) [REAL (rhs q___q), REAL v_v, REAL_REAL Q];
walther@59635
   245
    M__M = Rewrite ''Querkraft_Moment'' Q__Q;
walther@59635
   246
    M__M = SubProblem (''Biegelinie'', [''named'', ''integrate'', ''function''],
walther@59635
   247
      [''diff'', ''integration'', ''named'']) [REAL (rhs M__M), REAL v_v, REAL_REAL M_b];
walther@59635
   248
    N__N = (
walther@59637
   249
      (Rewrite ''Moment_Neigung'' ) #> (Rewrite ''make_fun_explicit'' )) M__M;
walther@59635
   250
    N__N = SubProblem (''Biegelinie'', [''named'', ''integrate'', ''function''],
walther@59635
   251
      [''diff'', ''integration'', ''named'']) [REAL (rhs N__N), REAL v_v, REAL_REAL id_abl];
walther@59635
   252
    B__B = SubProblem (''Biegelinie'', [''named'', ''integrate'', ''function''],
walther@59635
   253
      [''diff'', ''integration'', ''named'']) [REAL (rhs N__N), REAL v_v, REAL_REAL id_fun]
walther@59635
   254
  in
walther@59635
   255
    [Q__Q, M__M, N__N, B__B])"
wneuper@59472
   256
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   257
    [Specify.prep_met @{theory} "met_biege_ausbelast" [] Celem.e_metID
s1210629013@55373
   258
	    (["Biegelinien", "ausBelastung"],
wneuper@59540
   259
	      [("#Given" ,["Streckenlast q__q", "FunktionsVariable v_v",
wneuper@59540
   260
            "Biegelinie id_fun", "AbleitungBiegelinie id_abl"]),
wneuper@59540
   261
	       ("#Find"  ,["Funktionen fun_s"])],
s1210629013@55373
   262
	      {rew_ord'="tless_true", 
walther@59852
   263
	        rls' = Rule_Set.append_rules "erls_ausBelastung" Rule_Set.empty 
walther@59878
   264
				      [Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"),
walther@59871
   265
				        Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}),
walther@59871
   266
				        Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false})], 
s1210629013@55373
   267
				  calc = [], 
walther@59852
   268
				  srls = Rule_Set.append_rules "srls_ausBelastung" Rule_Set.empty 
walther@59878
   269
				      [Rule.Eval ("Prog_Expr.rhs", Prog_Expr.eval_rhs "eval_rhs_")], 
walther@59852
   270
				  prls = Rule_Set.empty, crls = Atools_erls, errpats = [], nrls = Rule_Set.empty},
wneuper@59551
   271
        @{thm belastung_zu_biegelinie.simps})]
wneuper@59472
   272
\<close>
wneuper@59429
   273
subsection \<open>Substitute the constraints into the equations\<close>
wneuper@59429
   274
wneuper@59504
   275
partial_function (tailrec) setzte_randbedingungen :: "bool list \<Rightarrow> bool list \<Rightarrow> bool list"
wneuper@59504
   276
  where 
walther@59635
   277
"setzte_randbedingungen fun_s r_b = (
walther@59635
   278
  let
walther@59635
   279
    b_1 = NTH 1 r_b;
walther@59635
   280
    f_s = filter_sameFunId (lhs b_1) fun_s;
walther@59635
   281
    e_1 = SubProblem (''Biegelinie'', [''makeFunctionTo'', ''equation''],
walther@59635
   282
            [''Equation'', ''fromFunction'']) [BOOL (hd f_s), BOOL b_1];
walther@59635
   283
    b_2 = NTH 2 r_b;
walther@59635
   284
    f_s = filter_sameFunId (lhs b_2) fun_s;
walther@59635
   285
    e_2 = SubProblem (''Biegelinie'', [''makeFunctionTo'', ''equation''],
walther@59635
   286
             [''Equation'', ''fromFunction'']) [BOOL (hd f_s), BOOL b_2];
walther@59635
   287
    b_3 = NTH 3 r_b;
walther@59635
   288
    f_s = filter_sameFunId (lhs b_3) fun_s;
walther@59635
   289
    e_3 = SubProblem (''Biegelinie'', [''makeFunctionTo'', ''equation''],
walther@59635
   290
             [''Equation'', ''fromFunction'']) [BOOL (hd f_s), BOOL b_3];
walther@59635
   291
    b_4 = NTH 4 r_b;
walther@59635
   292
    f_s = filter_sameFunId (lhs b_4) fun_s;
walther@59635
   293
    e_4 = SubProblem (''Biegelinie'', [''makeFunctionTo'', ''equation''],
walther@59635
   294
             [''Equation'', ''fromFunction'']) [BOOL (hd f_s), BOOL b_4]
walther@59635
   295
  in
walther@59635
   296
    [e_1, e_2, e_3, e_4])"
wneuper@59472
   297
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   298
    [Specify.prep_met @{theory} "met_biege_setzrand" [] Celem.e_metID
s1210629013@55373
   299
	    (["Biegelinien", "setzeRandbedingungenEin"],
s1210629013@55373
   300
	      [("#Given" , ["Funktionen fun_s", "Randbedingungen r_b"]),
s1210629013@55373
   301
	        ("#Find"  , ["Gleichungen equs'''"])],
walther@59852
   302
	      {rew_ord'="tless_true", rls'=Rule_Set.Empty, calc = [], srls = srls2, prls=Rule_Set.empty, crls = Atools_erls,
walther@59852
   303
          errpats = [], nrls = Rule_Set.empty},
wneuper@59551
   304
        @{thm setzte_randbedingungen.simps})]
wneuper@59472
   305
\<close>
wneuper@59429
   306
subsection \<open>Transform an equality into a function\<close>
wneuper@59429
   307
wneuper@59504
   308
        (*(M_b x = c_2 + c * x + -1 * q_0 / 2 * x ^^^ 2) (M_b L = 0) -->
wneuper@59504
   309
               0 = c_2 + c * L + -1 * q_0 / 2 * L ^^^ 2*)
wneuper@59504
   310
partial_function (tailrec) function_to_equality :: "bool \<Rightarrow> bool \<Rightarrow> bool"
wneuper@59504
   311
  where
walther@59635
   312
"function_to_equality fu_n su_b = (
walther@59635
   313
  let
walther@59635
   314
    fu_n = Take fu_n;
walther@59635
   315
    bd_v = argument_in (lhs fu_n);
walther@59635
   316
    va_l = argument_in (lhs su_b);
walther@59635
   317
    eq_u = Substitute [bd_v = va_l] fu_n; \<comment> \<open>([1], Res), M_b L = c_2 + c * L + -1 * q_0 / 2 * L ^^^ 2\<close>
walther@59635
   318
    eq_u = Substitute [su_b] eq_u
walther@59635
   319
  in
walther@59635
   320
    (Rewrite_Set ''norm_Rational'') eq_u)"
wneuper@59472
   321
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   322
    [Specify.prep_met @{theory} "met_equ_fromfun" [] Celem.e_metID
s1210629013@55373
   323
	    (["Equation","fromFunction"],
s1210629013@55373
   324
	      [("#Given" ,["functionEq fu_n","substitution su_b"]),
s1210629013@55373
   325
	        ("#Find"  ,["equality equ'''"])],
walther@59851
   326
	      {rew_ord'="tless_true", rls'=Rule_Set.Empty, calc = [],
walther@59852
   327
          srls = Rule_Set.append_rules "srls_in_EquationfromFunc" Rule_Set.empty
walther@59878
   328
				      [Rule.Eval("Prog_Expr.lhs", Prog_Expr.eval_lhs "eval_lhs_"),
walther@59878
   329
				        Rule.Eval("Prog_Expr.argument'_in", Prog_Expr.eval_argument_in "Prog_Expr.argument'_in")], 
walther@59852
   330
				  prls=Rule_Set.empty, crls = Atools_erls, errpats = [], nrls = Rule_Set.empty},
s1210629013@55373
   331
        (*(M_b x = c_2 + c * x + -1 * q_0 / 2 * x ^^^ 2) (M_b L = 0) -->
s1210629013@55373
   332
               0 = c_2 + c * L + -1 * q_0 / 2 * L ^^^ 2*)
wneuper@59551
   333
        @{thm function_to_equality.simps})]
wneuper@59472
   334
\<close>
wneuper@59549
   335
ML \<open>
wneuper@59549
   336
\<close> ML \<open>
wneuper@59549
   337
\<close>
neuper@37906
   338
end
neuper@37906
   339